doc-next-gen

Mathlib.RingTheory.NonUnitalSubring.Basic

Module docstring

{"# NonUnitalSubrings

Let R be a non-unital ring. We prove that non-unital subrings are a complete lattice, and that you can map (pushforward) and comap (pull back) them along ring homomorphisms.

We define the closure construction from Set R to NonUnitalSubring R, sending a subset of R to the non-unital subring it generates, and prove that it is a Galois insertion.

Main definitions

Notation used here:

(R : Type u) [NonUnitalRing R] (S : Type u) [NonUnitalRing S] (f g : R →ₙ+* S) (A : NonUnitalSubring R) (B : NonUnitalSubring S) (s : Set R)

  • instance : CompleteLattice (NonUnitalSubring R) : the complete lattice structure on the non-unital subrings.

  • NonUnitalSubring.center : the center of a non-unital ring R.

  • NonUnitalSubring.closure : non-unital subring closure of a set, i.e., the smallest non-unital subring that includes the set.

  • NonUnitalSubring.gi : closure : Set M → NonUnitalSubring M and coercion coe : NonUnitalSubring M → Set M form a GaloisInsertion.

  • comap f B : NonUnitalSubring A : the preimage of a non-unital subring B along the non-unital ring homomorphism f

  • map f A : NonUnitalSubring B : the image of a non-unital subring A along the non-unital ring homomorphism f.

  • Prod A B : NonUnitalSubring (R × S) : the product of non-unital subrings

  • f.range : NonUnitalSubring B : the range of the non-unital ring homomorphism f.

  • eq_locus f g : NonUnitalSubring R : given non-unital ring homomorphisms f g : R →ₙ+* S, the non-unital subring of R where f x = g x

Implementation notes

A non-unital subring is implemented as a NonUnitalSubsemiring which is also an additive subgroup.

Lattice inclusion (e.g. and ) is used rather than set notation ( and ), although is defined as membership of a non-unital subring's underlying set.

Tags

non-unital subring ","## top ","## comap ","## map ","## range ","## bot ","## inf ","## Center of a ring ","## NonUnitalSubring closure of a subset "}

NonUnitalSubring.list_sum_mem theorem
{l : List R} : (∀ x ∈ l, x ∈ s) → l.sum ∈ s
Full source
/-- Sum of a list of elements in a non-unital subring is in the non-unital subring. -/
protected theorem list_sum_mem {l : List R} : (∀ x ∈ l, x ∈ s) → l.sum ∈ s :=
  list_sum_mem
Closure of Non-Unital Subring under List Summation
Informal description
Let $R$ be a non-unital non-associative ring and $s$ a non-unital subring of $R$. For any list $l$ of elements of $R$, if every element $x \in l$ belongs to $s$, then the sum of all elements in $l$ also belongs to $s$.
NonUnitalSubring.multiset_sum_mem theorem
{R} [NonUnitalNonAssocRing R] (s : NonUnitalSubring R) (m : Multiset R) : (∀ a ∈ m, a ∈ s) → m.sum ∈ s
Full source
/-- Sum of a multiset of elements in a `NonUnitalSubring` of a `NonUnitalRing` is
in the `NonUnitalSubring`. -/
protected theorem multiset_sum_mem {R} [NonUnitalNonAssocRing R] (s : NonUnitalSubring R)
    (m : Multiset R) : (∀ a ∈ m, a ∈ s) → m.sum ∈ s :=
  multiset_sum_mem _
Closure of Non-Unital Subring under Multiset Summation
Informal description
Let $R$ be a non-unital non-associative ring and $s$ a non-unital subring of $R$. For any multiset $m$ of elements of $R$, if every element of $m$ belongs to $s$, then the sum of all elements in $m$ also belongs to $s$.
NonUnitalSubring.sum_mem theorem
{R : Type*} [NonUnitalNonAssocRing R] (s : NonUnitalSubring R) {ι : Type*} {t : Finset ι} {f : ι → R} (h : ∀ c ∈ t, f c ∈ s) : (∑ i ∈ t, f i) ∈ s
Full source
/-- Sum of elements in a `NonUnitalSubring` of a `NonUnitalRing` indexed by a `Finset`
is in the `NonUnitalSubring`. -/
protected theorem sum_mem {R : Type*} [NonUnitalNonAssocRing R] (s : NonUnitalSubring R)
    {ι : Type*} {t : Finset ι} {f : ι → R} (h : ∀ c ∈ t, f c ∈ s) : (∑ i ∈ t, f i) ∈ s :=
  sum_mem h
Sum of Elements in Non-Unital Subring is Closed Under Finite Summation
Informal description
Let $R$ be a non-unital non-associative ring and $s$ a non-unital subring of $R$. For any finite index set $\iota$, finite subset $t \subseteq \iota$, and function $f : \iota \to R$, if $f(c) \in s$ for all $c \in t$, then the sum $\sum_{i \in t} f(i)$ belongs to $s$.
NonUnitalSubring.mem_top theorem
(x : R) : x ∈ (⊤ : NonUnitalSubring R)
Full source
@[simp]
theorem mem_top (x : R) : x ∈ (⊤ : NonUnitalSubring R) :=
  Set.mem_univ x
Membership in Top Non-Unital Subring
Informal description
For any element $x$ in a non-unital ring $R$, $x$ belongs to the top non-unital subring of $R$ (which is $R$ itself).
NonUnitalSubring.coe_top theorem
: ((⊤ : NonUnitalSubring R) : Set R) = Set.univ
Full source
@[simp]
theorem coe_top : (( : NonUnitalSubring R) : Set R) = Set.univ :=
  rfl
Top Non-Unital Subring as Universal Set
Informal description
The underlying set of the top non-unital subring of a non-unital ring $R$ is equal to the universal set of $R$, i.e., $(\top : \text{NonUnitalSubring } R) = \text{Set.univ}$.
NonUnitalSubring.topEquiv definition
: (⊤ : NonUnitalSubring R) ≃+* R
Full source
/-- The ring equiv between the top element of `NonUnitalSubring R` and `R`. -/
@[simps!]
def topEquiv : (⊤ : NonUnitalSubring R) ≃+* R := NonUnitalSubsemiring.topEquiv
Ring isomorphism between top non-unital subring and the ring itself
Informal description
The ring isomorphism between the top element of the lattice of non-unital subrings of $R$ (which is $R$ itself) and $R$. This equivalence preserves both the additive and multiplicative structures.
NonUnitalSubring.comap definition
{F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] (f : F) (s : NonUnitalSubring S) : NonUnitalSubring R
Full source
/-- The preimage of a `NonUnitalSubring` along a ring homomorphism is a `NonUnitalSubring`. -/
def comap {F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S]
    [FunLike F R S] [NonUnitalRingHomClass F R S] (f : F) (s : NonUnitalSubring S) :
    NonUnitalSubring R :=
  { s.toSubsemigroup.comap (f : R →ₙ* S), s.toAddSubgroup.comap (f : R →+ S) with
    carrier := f ⁻¹' s.carrier }
Preimage of a non-unital subring under a ring homomorphism
Informal description
Given a non-unital ring homomorphism $f : R \to S$ and a non-unital subring $s$ of $S$, the preimage of $s$ under $f$ forms a non-unital subring of $R$. Specifically, the carrier set of this subring is $f^{-1}(s)$, and it inherits the additive and multiplicative substructures from $R$ via the homomorphism $f$.
NonUnitalSubring.coe_comap theorem
(s : NonUnitalSubring S) (f : F) : (s.comap f : Set R) = f ⁻¹' s
Full source
@[simp]
theorem coe_comap (s : NonUnitalSubring S) (f : F) : (s.comap f : Set R) = f ⁻¹' s :=
  rfl
Preimage of Non-unital Subring Under Homomorphism Equals Set Preimage
Informal description
For any non-unital subring $s$ of a non-unital ring $S$ and any non-unital ring homomorphism $f : R \to S$, the underlying set of the preimage subring $\text{comap}(f, s)$ is equal to the preimage of $s$ under $f$, i.e., $f^{-1}(s)$.
NonUnitalSubring.mem_comap theorem
{s : NonUnitalSubring S} {f : F} {x : R} : x ∈ s.comap f ↔ f x ∈ s
Full source
@[simp]
theorem mem_comap {s : NonUnitalSubring S} {f : F} {x : R} : x ∈ s.comap fx ∈ s.comap f ↔ f x ∈ s :=
  Iff.rfl
Characterization of Membership in Preimage Subring
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f : R \to S$ be a non-unital ring homomorphism. For any non-unital subring $s$ of $S$ and any element $x \in R$, we have $x$ belongs to the preimage subring $s.\text{comap}\,f$ if and only if $f(x)$ belongs to $s$.
NonUnitalSubring.comap_comap theorem
(s : NonUnitalSubring T) (g : S →ₙ+* T) (f : R →ₙ+* S) : (s.comap g).comap f = s.comap (g.comp f)
Full source
theorem comap_comap (s : NonUnitalSubring T) (g : S →ₙ+* T) (f : R →ₙ+* S) :
    (s.comap g).comap f = s.comap (g.comp f) :=
  rfl
Composition of Preimages of Non-unital Subrings under Ring Homomorphisms
Informal description
Let $R$, $S$, and $T$ be non-unital non-associative rings, and let $f \colon R \to S$ and $g \colon S \to T$ be non-unital ring homomorphisms. For any non-unital subring $s$ of $T$, the preimage of $s$ under $g$ composed with $f$ is equal to the preimage of $s$ under the composition $g \circ f$. In other words, $(s.comap\ g).comap\ f = s.comap\ (g \circ f)$.
NonUnitalSubring.map definition
{F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] (f : F) (s : NonUnitalSubring R) : NonUnitalSubring S
Full source
/-- The image of a `NonUnitalSubring` along a ring homomorphism is a `NonUnitalSubring`. -/
def map {F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S]
    [FunLike F R S] [NonUnitalRingHomClass F R S] (f : F) (s : NonUnitalSubring R) :
    NonUnitalSubring S :=
  { s.toSubsemigroup.map (f : R →ₙ* S), s.toAddSubgroup.map (f : R →+ S) with
    carrier := f '' s.carrier }
Image of a non-unital subring under a ring homomorphism
Informal description
Given a non-unital ring homomorphism $f \colon R \to S$ between non-unital non-associative rings $R$ and $S$, and a non-unital subring $s$ of $R$, the image of $s$ under $f$ is a non-unital subring of $S$. The underlying set of this subring is the image of the underlying set of $s$ under $f$.
NonUnitalSubring.coe_map theorem
(f : F) (s : NonUnitalSubring R) : (s.map f : Set S) = f '' s
Full source
@[simp]
theorem coe_map (f : F) (s : NonUnitalSubring R) : (s.map f : Set S) = f '' s :=
  rfl
Image of Non-unital Subring Under Homomorphism Equals Image of Underlying Set
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be a non-unital ring homomorphism. For any non-unital subring $s$ of $R$, the underlying set of the image subring $s.map(f)$ is equal to the image of the underlying set of $s$ under $f$, i.e., $s.map(f) = f(s)$ as subsets of $S$.
NonUnitalSubring.mem_map theorem
{f : F} {s : NonUnitalSubring R} {y : S} : y ∈ s.map f ↔ ∃ x ∈ s, f x = y
Full source
@[simp]
theorem mem_map {f : F} {s : NonUnitalSubring R} {y : S} : y ∈ s.map fy ∈ s.map f ↔ ∃ x ∈ s, f x = y :=
  Set.mem_image _ _ _
Characterization of Elements in the Image of a Non-unital Subring under a Homomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be a non-unital ring homomorphism. For any non-unital subring $s$ of $R$ and any element $y \in S$, we have $y \in f(s)$ if and only if there exists $x \in s$ such that $f(x) = y$.
NonUnitalSubring.map_id theorem
: s.map (NonUnitalRingHom.id R) = s
Full source
@[simp]
theorem map_id : s.map (NonUnitalRingHom.id R) = s :=
  SetLike.coe_injective <| Set.image_id _
Image of Subring under Identity Homomorphism is Itself
Informal description
For any non-unital subring $s$ of a non-unital non-associative ring $R$, the image of $s$ under the identity ring homomorphism $\text{id}_R \colon R \to R$ is equal to $s$ itself, i.e., $s.\text{map}(\text{id}_R) = s$.
NonUnitalSubring.map_map theorem
(g : S →ₙ+* T) (f : R →ₙ+* S) : (s.map f).map g = s.map (g.comp f)
Full source
theorem map_map (g : S →ₙ+* T) (f : R →ₙ+* S) : (s.map f).map g = s.map (g.comp f) :=
  SetLike.coe_injective <| Set.image_image _ _ _
Composition of Images of Non-unital Subrings under Ring Homomorphisms
Informal description
Let $R$, $S$, and $T$ be non-unital non-associative rings, and let $f \colon R \to S$ and $g \colon S \to T$ be non-unital ring homomorphisms. For any non-unital subring $s$ of $R$, the image of $s$ under the composition $g \circ f$ is equal to the image of the image of $s$ under $f$ under $g$, i.e., $(s.\text{map}(f)).\text{map}(g) = s.\text{map}(g \circ f)$.
NonUnitalSubring.map_le_iff_le_comap theorem
{f : F} {s : NonUnitalSubring R} {t : NonUnitalSubring S} : s.map f ≤ t ↔ s ≤ t.comap f
Full source
theorem map_le_iff_le_comap {f : F} {s : NonUnitalSubring R} {t : NonUnitalSubring S} :
    s.map f ≤ t ↔ s ≤ t.comap f :=
  Set.image_subset_iff
Image-Preimage Inclusion for Non-Unital Subrings
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be a non-unital ring homomorphism. For any non-unital subring $s$ of $R$ and any non-unital subring $t$ of $S$, the image of $s$ under $f$ is contained in $t$ if and only if $s$ is contained in the preimage of $t$ under $f$. In other words, $f(s) \subseteq t \leftrightarrow s \subseteq f^{-1}(t)$.
NonUnitalSubring.gc_map_comap theorem
(f : F) : GaloisConnection (map f : NonUnitalSubring R → NonUnitalSubring S) (comap f)
Full source
theorem gc_map_comap (f : F) :
    GaloisConnection (map f : NonUnitalSubring R → NonUnitalSubring S) (comap f) := fun _S _T =>
  map_le_iff_le_comap
Galois Connection Between Image and Preimage of Non-Unital Subrings Under a Ring Homomorphism
Informal description
For any non-unital ring homomorphism $f \colon R \to S$, the pair of functions $\text{map}(f) \colon \text{NonUnitalSubring}(R) \to \text{NonUnitalSubring}(S)$ and $\text{comap}(f) \colon \text{NonUnitalSubring}(S) \to \text{NonUnitalSubring}(R)$ form a Galois connection. That is, for any non-unital subring $A$ of $R$ and any non-unital subring $B$ of $S$, we have $\text{map}(f)(A) \leq B$ if and only if $A \leq \text{comap}(f)(B)$.
NonUnitalSubring.equivMapOfInjective definition
(f : F) (hf : Function.Injective (f : R → S)) : s ≃+* s.map f
Full source
/-- A `NonUnitalSubring` is isomorphic to its image under an injective function -/
noncomputable def equivMapOfInjective (f : F) (hf : Function.Injective (f : R → S)) :
    s ≃+* s.map f :=
  {
    Equiv.Set.image f s
      hf with
    map_mul' := fun _ _ => Subtype.ext (map_mul f _ _)
    map_add' := fun _ _ => Subtype.ext (map_add f _ _) }
Isomorphism between a non-unital subring and its image under an injective homomorphism
Informal description
Given an injective non-unital ring homomorphism $f \colon R \to S$ and a non-unital subring $s$ of $R$, there exists a ring isomorphism between $s$ and its image $f(s)$ under $f$. This isomorphism preserves both the additive and multiplicative structures.
NonUnitalSubring.coe_equivMapOfInjective_apply theorem
(f : F) (hf : Function.Injective f) (x : s) : (equivMapOfInjective s f hf x : S) = f x
Full source
@[simp]
theorem coe_equivMapOfInjective_apply (f : F) (hf : Function.Injective f) (x : s) :
    (equivMapOfInjective s f hf x : S) = f x :=
  rfl
Image of Element under Isomorphism Induced by Injective Homomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be an injective non-unital ring homomorphism. For any non-unital subring $s$ of $R$ and any element $x \in s$, the image of $x$ under the isomorphism $\text{equivMapOfInjective}(s, f, hf)$ equals $f(x)$ when viewed as an element of $S$.
NonUnitalRingHom.range definition
{R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] (f : R →ₙ+* S) : NonUnitalSubring S
Full source
/-- The range of a ring homomorphism, as a `NonUnitalSubring` of the target.
See Note [range copy pattern]. -/
def range {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S]
    (f : R →ₙ+* S) : NonUnitalSubring S :=
  (( : NonUnitalSubring R).map f).copy (Set.range f) Set.image_univ.symm
Range of a non-unital ring homomorphism
Informal description
Given a non-unital ring homomorphism $f \colon R \to S$ between non-unital non-associative rings $R$ and $S$, the range of $f$ is the non-unital subring of $S$ consisting of all elements $y \in S$ such that there exists $x \in R$ with $f(x) = y$. This is constructed as the image of the top non-unital subring of $R$ under $f$, with the underlying set being the range of $f$ as a function.
NonUnitalRingHom.coe_range theorem
: (f.range : Set S) = Set.range f
Full source
@[simp]
theorem coe_range : (f.range : Set S) = Set.range f :=
  rfl
Range of Non-unital Ring Homomorphism as Set
Informal description
For a non-unital ring homomorphism $f \colon R \to S$, the underlying set of the range of $f$ (as a non-unital subring of $S$) is equal to the range of $f$ as a function, i.e., $\{f(x) \mid x \in R\}$.
NonUnitalRingHom.mem_range theorem
{f : R →ₙ+* S} {y : S} : y ∈ f.range ↔ ∃ x, f x = y
Full source
@[simp]
theorem mem_range {f : R →ₙ+* S} {y : S} : y ∈ f.rangey ∈ f.range ↔ ∃ x, f x = y :=
  Iff.rfl
Characterization of Elements in the Range of a Non-Unital Ring Homomorphism
Informal description
For a non-unital ring homomorphism $f \colon R \to S$ and an element $y \in S$, $y$ belongs to the range of $f$ if and only if there exists an element $x \in R$ such that $f(x) = y$.
NonUnitalRingHom.range_eq_map theorem
(f : R →ₙ+* S) : f.range = NonUnitalSubring.map f ⊤
Full source
theorem range_eq_map (f : R →ₙ+* S) : f.range = NonUnitalSubring.map f  := by ext; simp
Range of a Non-Unital Ring Homomorphism as Image of Top Subring
Informal description
For any non-unital ring homomorphism $f \colon R \to S$ between non-unital non-associative rings $R$ and $S$, the range of $f$ is equal to the image of the top non-unital subring of $R$ under $f$.
NonUnitalRingHom.mem_range_self theorem
(f : R →ₙ+* S) (x : R) : f x ∈ f.range
Full source
theorem mem_range_self (f : R →ₙ+* S) (x : R) : f x ∈ f.range :=
  mem_range.mpr ⟨x, rfl⟩
Image of Element under Non-Unital Ring Homomorphism is in Range
Informal description
For any non-unital ring homomorphism $f \colon R \to S$ and any element $x \in R$, the image $f(x)$ belongs to the range of $f$.
NonUnitalRingHom.map_range theorem
: f.range.map g = (g.comp f).range
Full source
theorem map_range : f.range.map g = (g.comp f).range := by
  simpa only [range_eq_map] using ( : NonUnitalSubring R).map_map g f
Image of Range under Composition of Non-unital Ring Homomorphisms
Informal description
Let $R$, $S$, and $T$ be non-unital non-associative rings, and let $f \colon R \to S$ and $g \colon S \to T$ be non-unital ring homomorphisms. Then the image of the range of $f$ under $g$ is equal to the range of the composition $g \circ f$.
NonUnitalRingHom.fintypeRange instance
[Fintype R] [DecidableEq S] (f : R →ₙ+* S) : Fintype (range f)
Full source
/-- The range of a ring homomorphism is a fintype, if the domain is a fintype.
Note: this instance can form a diamond with `Subtype.fintype` in the
  presence of `Fintype S`. -/
instance fintypeRange [Fintype R] [DecidableEq S] (f : R →ₙ+* S) : Fintype (range f) :=
  Set.fintypeRange f
Finiteness of the Range of a Non-unital Ring Homomorphism
Informal description
For any non-unital ring homomorphism $f \colon R \to S$ between non-unital non-associative rings, if $R$ is a finite type and equality is decidable on $S$, then the range of $f$ is also a finite type.
NonUnitalSubring.instBot instance
: Bot (NonUnitalSubring R)
Full source
instance : Bot (NonUnitalSubring R) :=
  ⟨(0 : R →ₙ+* R).range⟩
Existence of Trivial Non-Unital Subring
Informal description
The collection of all non-unital subrings of a non-unital ring $R$ has a bottom element, which is the trivial subring $\{0\}$.
NonUnitalSubring.instInhabited instance
: Inhabited (NonUnitalSubring R)
Full source
instance : Inhabited (NonUnitalSubring R) :=
  ⟨⊥⟩
Non-Unital Subrings are Inhabited
Informal description
For any non-unital ring $R$, the collection of non-unital subrings of $R$ is inhabited (i.e., contains at least one element).
NonUnitalSubring.mem_bot theorem
{x : R} : x ∈ (⊥ : NonUnitalSubring R) ↔ x = 0
Full source
theorem mem_bot {x : R} : x ∈ (⊥ : NonUnitalSubring R)x ∈ (⊥ : NonUnitalSubring R) ↔ x = 0 :=
  show x ∈ ((⊥ : NonUnitalSubring R) : Set R)x ∈ ((⊥ : NonUnitalSubring R) : Set R) ↔ x = 0 by rw [coe_bot, Set.mem_singleton_iff]
Membership in Trivial Non-Unital Subring: $x \in \{0\} \leftrightarrow x = 0$
Informal description
For any element $x$ in a non-unital ring $R$, $x$ belongs to the trivial non-unital subring $\{0\}$ if and only if $x = 0$.
NonUnitalSubring.coe_inf theorem
(p p' : NonUnitalSubring R) : ((p ⊓ p' : NonUnitalSubring R) : Set R) = (p : Set R) ∩ p'
Full source
@[simp]
theorem coe_inf (p p' : NonUnitalSubring R) :
    ((p ⊓ p' : NonUnitalSubring R) : Set R) = (p : Set R) ∩ p' :=
  rfl
Infimum of Non-Unital Subrings as Intersection
Informal description
For any two non-unital subrings $p$ and $p'$ of a non-unital non-associative ring $R$, the underlying set of their infimum $p \sqcap p'$ is equal to the intersection of the underlying sets of $p$ and $p'$, i.e., $(p \sqcap p') = p \cap p'$.
NonUnitalSubring.mem_inf theorem
{p p' : NonUnitalSubring R} {x : R} : x ∈ p ⊓ p' ↔ x ∈ p ∧ x ∈ p'
Full source
@[simp]
theorem mem_inf {p p' : NonUnitalSubring R} {x : R} : x ∈ p ⊓ p'x ∈ p ⊓ p' ↔ x ∈ p ∧ x ∈ p' :=
  Iff.rfl
Membership in Intersection of Non-unital Subrings: $x \in p \sqcap p' \leftrightarrow x \in p \land x \in p'$
Informal description
For any two non-unital subrings $p$ and $p'$ of a non-unital non-associative ring $R$, an element $x \in R$ belongs to the infimum $p \sqcap p'$ if and only if $x$ belongs to both $p$ and $p'$.
NonUnitalSubring.instInfSet instance
: InfSet (NonUnitalSubring R)
Full source
instance : InfSet (NonUnitalSubring R) :=
  ⟨fun s =>
    NonUnitalSubring.mk' (⋂ t ∈ s, ↑t) (⨅ t ∈ s, NonUnitalSubring.toSubsemigroup t)
      (⨅ t ∈ s, NonUnitalSubring.toAddSubgroup t) (by simp) (by simp)⟩
Complete Lattice Structure on Non-Unital Subrings
Informal description
The collection of all non-unital subrings of a non-unital non-associative ring $R$ forms a complete lattice with respect to inclusion, where the infimum (meet) of a family of non-unital subrings is given by their intersection.
NonUnitalSubring.coe_sInf theorem
(S : Set (NonUnitalSubring R)) : ((sInf S : NonUnitalSubring R) : Set R) = ⋂ s ∈ S, ↑s
Full source
@[simp, norm_cast]
theorem coe_sInf (S : Set (NonUnitalSubring R)) :
    ((sInf S : NonUnitalSubring R) : Set R) = ⋂ s ∈ S, ↑s :=
  rfl
Infimum of Non-Unital Subrings as Intersection
Informal description
For any collection $S$ of non-unital subrings of a non-unital non-associative ring $R$, the underlying set of the infimum of $S$ is equal to the intersection of the underlying sets of all subrings in $S$. That is, $\big(\inf S\big) = \bigcap_{s \in S} s$.
NonUnitalSubring.mem_sInf theorem
{S : Set (NonUnitalSubring R)} {x : R} : x ∈ sInf S ↔ ∀ p ∈ S, x ∈ p
Full source
theorem mem_sInf {S : Set (NonUnitalSubring R)} {x : R} : x ∈ sInf Sx ∈ sInf S ↔ ∀ p ∈ S, x ∈ p :=
  Set.mem_iInter₂
Characterization of Membership in the Infimum of Non-Unital Subrings
Informal description
For any family of non-unital subrings $S$ of a non-unital ring $R$ and any element $x \in R$, $x$ belongs to the infimum of $S$ (i.e., the intersection of all subrings in $S$) if and only if $x$ belongs to every subring in $S$. In symbols: \[ x \in \bigwedge S \leftrightarrow \forall p \in S, x \in p. \]
NonUnitalSubring.coe_iInf theorem
{ι : Sort*} {S : ι → NonUnitalSubring R} : (↑(⨅ i, S i) : Set R) = ⋂ i, S i
Full source
@[simp, norm_cast]
theorem coe_iInf {ι : Sort*} {S : ι → NonUnitalSubring R} : (↑(⨅ i, S i) : Set R) = ⋂ i, S i := by
  simp only [iInf, coe_sInf, Set.biInter_range]
Infimum of Non-Unital Subrings as Intersection
Informal description
For any family of non-unital subrings $\{S_i\}_{i \in \iota}$ of a non-unital non-associative ring $R$, the underlying set of their infimum $\bigsqcap_i S_i$ is equal to the intersection $\bigcap_i S_i$ of their underlying sets.
NonUnitalSubring.mem_iInf theorem
{ι : Sort*} {S : ι → NonUnitalSubring R} {x : R} : (x ∈ ⨅ i, S i) ↔ ∀ i, x ∈ S i
Full source
theorem mem_iInf {ι : Sort*} {S : ι → NonUnitalSubring R} {x : R} :
    (x ∈ ⨅ i, S i) ↔ ∀ i, x ∈ S i := by simp only [iInf, mem_sInf, Set.forall_mem_range]
Characterization of Membership in Infimum of Non-Unital Subrings
Informal description
Let $R$ be a non-unital non-associative ring, and let $\{S_i\}_{i \in \iota}$ be a family of non-unital subrings of $R$. An element $x \in R$ belongs to the infimum $\bigsqcap_i S_i$ of the family $\{S_i\}$ if and only if $x$ belongs to $S_i$ for every index $i \in \iota$.
NonUnitalSubring.sInf_toSubsemigroup theorem
(s : Set (NonUnitalSubring R)) : (sInf s).toSubsemigroup = ⨅ t ∈ s, NonUnitalSubring.toSubsemigroup t
Full source
@[simp]
theorem sInf_toSubsemigroup (s : Set (NonUnitalSubring R)) :
    (sInf s).toSubsemigroup = ⨅ t ∈ s, NonUnitalSubring.toSubsemigroup t :=
  mk'_toSubsemigroup _ _
Infimum of Non-Unital Subrings Preserves Underlying Multiplicative Subsemigroup
Informal description
For any collection $s$ of non-unital subrings of a non-unital non-associative ring $R$, the underlying multiplicative subsemigroup of the infimum of $s$ is equal to the infimum of the underlying multiplicative subsemigroups of the subrings in $s$.
NonUnitalSubring.sInf_toAddSubgroup theorem
(s : Set (NonUnitalSubring R)) : (sInf s).toAddSubgroup = ⨅ t ∈ s, NonUnitalSubring.toAddSubgroup t
Full source
@[simp]
theorem sInf_toAddSubgroup (s : Set (NonUnitalSubring R)) :
    (sInf s).toAddSubgroup = ⨅ t ∈ s, NonUnitalSubring.toAddSubgroup t :=
  mk'_toAddSubgroup _ _
Infimum of Non-unital Subrings Preserves Underlying Additive Subgroup
Informal description
For any collection $s$ of non-unital subrings of a non-unital non-associative ring $R$, the underlying additive subgroup of the infimum of $s$ is equal to the infimum of the underlying additive subgroups of the subrings in $s$. In other words, $(\inf s).\text{toAddSubgroup} = \inf_{t \in s} t.\text{toAddSubgroup}$.
NonUnitalSubring.instCompleteLattice instance
: CompleteLattice (NonUnitalSubring R)
Full source
/-- `NonUnitalSubring`s of a ring form a complete lattice. -/
instance : CompleteLattice (NonUnitalSubring R) :=
  { completeLatticeOfInf (NonUnitalSubring R) fun _s =>
      IsGLB.of_image (@fun _ _ : NonUnitalSubring R => SetLike.coe_subset_coe)
        isGLB_biInf with
    bot := 
    bot_le := fun s _x hx => (mem_bot.mp hx).symmzero_mem s
    top := 
    le_top := fun _ _ _ => trivial
    inf := (· ⊓ ·)
    inf_le_left := fun _ _ _ => And.left
    inf_le_right := fun _ _ _ => And.right
    le_inf := fun _s _t₁ _t₂ h₁ h₂ _x hx => ⟨h₁ hx, h₂ hx⟩ }
Complete Lattice Structure on Non-Unital Subrings
Informal description
The collection of all non-unital subrings of a non-unital non-associative ring $R$ forms a complete lattice, where the order is given by inclusion, meets are intersections, and joins are the smallest non-unital subrings containing the unions.
NonUnitalSubring.center definition
: NonUnitalSubring R
Full source
/-- The center of a ring `R` is the set of elements that commute with everything in `R` -/
def center : NonUnitalSubring R :=
  { NonUnitalSubsemiring.center R with
    neg_mem' := Set.neg_mem_center }
Center of a non-unital ring
Informal description
The center of a non-unital ring $R$ is the subset of elements that commute with every element in $R$. It forms a non-unital subring of $R$ and is closed under negation.
NonUnitalSubring.coe_center theorem
: ↑(center R) = Set.center R
Full source
theorem coe_center : ↑(center R) = Set.center R :=
  rfl
Equality of Non-unital Subring Center and Set of Central Elements
Informal description
The underlying set of the center of a non-unital ring $R$ is equal to the set of all central elements in $R$, i.e., $\overline{\text{center}(R)} = \{z \in R \mid \forall x \in R, zx = xz\}$.
NonUnitalSubring.center_toNonUnitalSubsemiring theorem
: (center R).toNonUnitalSubsemiring = NonUnitalSubsemiring.center R
Full source
@[simp]
theorem center_toNonUnitalSubsemiring :
    (center R).toNonUnitalSubsemiring = NonUnitalSubsemiring.center R :=
  rfl
Equality of Non-unital Subring Center and Non-unital Subsemiring Center
Informal description
The underlying non-unital subsemiring of the center of a non-unital ring $R$ is equal to the center of $R$ considered as a non-unital subsemiring.
NonUnitalSubring.centerCongr definition
{S} [NonUnitalNonAssocRing S] (e : R ≃+* S) : center R ≃+* center S
Full source
/-- The center of isomorphic (not necessarily unital or associative) rings are isomorphic. -/
@[simps!] def centerCongr {S} [NonUnitalNonAssocRing S] (e : R ≃+* S) : centercenter R ≃+* center S :=
  NonUnitalSubsemiring.centerCongr e
Isomorphism of centers under ring isomorphism
Informal description
Given a non-unital non-associative ring isomorphism $e : R \simeq+* S$, the centers of $R$ and $S$ are isomorphic as non-unital rings. Specifically, the restriction of $e$ to the center of $R$ induces an isomorphism between the center of $R$ and the center of $S$.
NonUnitalSubring.centerToMulOpposite definition
: center R ≃+* center Rᵐᵒᵖ
Full source
/-- The center of a (not necessarily uintal or associative) ring
is isomorphic to the center of its opposite. -/
@[simps!] def centerToMulOpposite : centercenter R ≃+* center Rᵐᵒᵖ :=
  NonUnitalSubsemiring.centerToMulOpposite
Isomorphism between the center of a non-unital ring and its opposite
Informal description
The center of a non-unital ring $R$ is isomorphic as a non-unital ring to the center of its opposite ring $R^{\mathrm{op}}$. This isomorphism preserves the additive and multiplicative structures.
NonUnitalSubring.decidableMemCenter instance
[DecidableEq R] [Fintype R] : DecidablePred (· ∈ center R)
Full source
instance decidableMemCenter [DecidableEq R] [Fintype R] : DecidablePred (· ∈ center R) := fun _ =>
  decidable_of_iff' _ mem_center_iff
Decidability of Membership in the Center of a Finite Non-unital Ring
Informal description
For any non-unital ring $R$ with a decidable equality and a finite underlying set, membership in the center of $R$ is decidable. That is, there exists an algorithm to determine whether a given element $z \in R$ commutes with all elements of $R$.
NonUnitalSubring.center_eq_top theorem
(R) [NonUnitalCommRing R] : center R = ⊤
Full source
@[simp]
theorem center_eq_top (R) [NonUnitalCommRing R] : center R =  :=
  SetLike.coe_injective (Set.center_eq_univ R)
Center Equals Entire Ring in Non-unital Commutative Rings
Informal description
For a non-unital commutative ring $R$, the center of $R$ is equal to the entire ring, i.e., $\text{center}(R) = R$.
NonUnitalSubring.closure definition
(s : Set R) : NonUnitalSubring R
Full source
/-- The `NonUnitalSubring` generated by a set. -/
def closure (s : Set R) : NonUnitalSubring R :=
  sInf {S | s ⊆ S}
Non-unital subring generated by a set
Informal description
The `NonUnitalSubring.closure` of a subset $s$ of a non-unital non-associative ring $R$ is the smallest non-unital subring of $R$ that contains $s$. It is defined as the infimum of all non-unital subrings of $R$ that contain $s$.
NonUnitalSubring.mem_closure theorem
{x : R} {s : Set R} : x ∈ closure s ↔ ∀ S : NonUnitalSubring R, s ⊆ S → x ∈ S
Full source
theorem mem_closure {x : R} {s : Set R} : x ∈ closure sx ∈ closure s ↔ ∀ S : NonUnitalSubring R, s ⊆ S → x ∈ S :=
  mem_sInf
Characterization of Membership in Non-unital Subring Closure
Informal description
For any element $x$ in a non-unital non-associative ring $R$ and any subset $s$ of $R$, $x$ belongs to the non-unital subring closure of $s$ if and only if $x$ is contained in every non-unital subring $S$ of $R$ that contains $s$. In symbols: \[ x \in \text{closure}(s) \leftrightarrow \forall S \leq R, s \subseteq S \to x \in S. \]
NonUnitalSubring.subset_closure theorem
{s : Set R} : s ⊆ closure s
Full source
/-- The `NonUnitalSubring` generated by a set includes the set. -/
@[simp, aesop safe 20 apply (rule_sets := [SetLike])]
theorem subset_closure {s : Set R} : s ⊆ closure s := fun _x hx => mem_closure.2 fun _S hS => hS hx
Subset is Contained in its Non-unital Subring Closure
Informal description
For any subset $s$ of a non-unital non-associative ring $R$, the set $s$ is contained in the non-unital subring generated by $s$, i.e., $s \subseteq \text{closure}(s)$.
NonUnitalSubring.not_mem_of_not_mem_closure theorem
{s : Set R} {P : R} (hP : P ∉ closure s) : P ∉ s
Full source
theorem not_mem_of_not_mem_closure {s : Set R} {P : R} (hP : P ∉ closure s) : P ∉ s := fun h =>
  hP (subset_closure h)
Non-membership in Subring Closure Implies Non-membership in Original Set
Informal description
For any subset $s$ of a non-unital non-associative ring $R$ and any element $P \in R$, if $P$ does not belong to the non-unital subring closure of $s$, then $P$ does not belong to $s$.
NonUnitalSubring.closure_le theorem
{s : Set R} {t : NonUnitalSubring R} : closure s ≤ t ↔ s ⊆ t
Full source
/-- A `NonUnitalSubring` `t` includes `closure s` if and only if it includes `s`. -/
@[simp]
theorem closure_le {s : Set R} {t : NonUnitalSubring R} : closureclosure s ≤ t ↔ s ⊆ t :=
  ⟨Set.Subset.trans subset_closure, fun h => sInf_le h⟩
Characterization of Non-Unital Subring Closure Containment
Informal description
For any subset $s$ of a non-unital non-associative ring $R$ and any non-unital subring $t$ of $R$, the closure of $s$ is contained in $t$ if and only if $s$ is contained in $t$.
NonUnitalSubring.closure_mono theorem
⦃s t : Set R⦄ (h : s ⊆ t) : closure s ≤ closure t
Full source
/-- `NonUnitalSubring` closure of a set is monotone in its argument: if `s ⊆ t`,
then `closure s ≤ closure t`. -/
@[gcongr]
theorem closure_mono ⦃s t : Set R⦄ (h : s ⊆ t) : closure s ≤ closure t :=
  closure_le.2 <| Set.Subset.trans h subset_closure
Monotonicity of Non-Unital Subring Closure
Informal description
For any subsets $s$ and $t$ of a non-unital non-associative ring $R$, if $s \subseteq t$, then the non-unital subring generated by $s$ is contained in the non-unital subring generated by $t$, i.e., $\text{closure}(s) \leq \text{closure}(t)$.
NonUnitalSubring.closure_eq_of_le theorem
{s : Set R} {t : NonUnitalSubring R} (h₁ : s ⊆ t) (h₂ : t ≤ closure s) : closure s = t
Full source
theorem closure_eq_of_le {s : Set R} {t : NonUnitalSubring R} (h₁ : s ⊆ t) (h₂ : t ≤ closure s) :
    closure s = t :=
  le_antisymm (closure_le.2 h₁) h₂
Characterization of Non-Unital Subring Closure Equality
Informal description
For any subset $s$ of a non-unital non-associative ring $R$ and any non-unital subring $t$ of $R$, if $s \subseteq t$ and $t$ is contained in the closure of $s$, then the closure of $s$ equals $t$.
NonUnitalSubring.closure_induction theorem
{s : Set R} {p : (x : R) → x ∈ closure s → Prop} (mem : ∀ (x) (hx : x ∈ s), p x (subset_closure hx)) (zero : p 0 (zero_mem _)) (add : ∀ x y hx hy, p x hx → p y hy → p (x + y) (add_mem hx hy)) (neg : ∀ x hx, p x hx → p (-x) (neg_mem hx)) (mul : ∀ x y hx hy, p x hx → p y hy → p (x * y) (mul_mem hx hy)) {x} (hx : x ∈ closure s) : p x hx
Full source
/-- An induction principle for closure membership. If `p` holds for `0`, `1`, and all elements
of `s`, and is preserved under addition, negation, and multiplication, then `p` holds for all
elements of the closure of `s`. -/
@[elab_as_elim]
theorem closure_induction {s : Set R} {p : (x : R) → x ∈ closure s → Prop}
    (mem : ∀ (x) (hx : x ∈ s), p x (subset_closure hx)) (zero : p 0 (zero_mem _))
    (add : ∀ x y hx hy, p x hx → p y hy → p (x + y) (add_mem hx hy))
    (neg : ∀ x hx, p x hx → p (-x) (neg_mem hx))
    (mul : ∀ x y hx hy, p x hx → p y hy → p (x * y) (mul_mem hx hy))
    {x} (hx : x ∈ closure s) : p x hx :=
  let K : NonUnitalSubring R :=
    { carrier := { x | ∃ hx, p x hx }
      mul_mem' := fun ⟨_, hpx⟩ ⟨_, hpy⟩ ↦ ⟨_, mul _ _ _ _ hpx hpy⟩
      add_mem' := fun ⟨_, hpx⟩ ⟨_, hpy⟩ ↦ ⟨_, add _ _ _ _ hpx hpy⟩
      neg_mem' := fun ⟨_, hpx⟩ ↦ ⟨_, neg _ _ hpx⟩
      zero_mem' := ⟨_, zero⟩ }
  closure_le (t := K) |>.mpr (fun y hy ↦ ⟨subset_closure hy, mem y hy⟩) hx |>.elim fun _ ↦ id
Induction Principle for Non-Unital Subring Closure Membership
Informal description
Let $R$ be a non-unital non-associative ring and $s$ a subset of $R$. For a predicate $p$ on elements of the non-unital subring closure of $s$, if: 1. $p(x)$ holds for all $x \in s$, 2. $p(0)$ holds, 3. $p$ is preserved under addition (i.e., $p(x)$ and $p(y)$ imply $p(x + y)$), 4. $p$ is preserved under negation (i.e., $p(x)$ implies $p(-x)$), 5. $p$ is preserved under multiplication (i.e., $p(x)$ and $p(y)$ imply $p(x \cdot y)$), then $p(x)$ holds for all $x$ in the non-unital subring closure of $s$.
NonUnitalSubring.closure_induction₂ theorem
{s : Set R} {p : (x y : R) → x ∈ closure s → y ∈ closure s → Prop} (mem_mem : ∀ (x) (y) (hx : x ∈ s) (hy : y ∈ s), p x y (subset_closure hx) (subset_closure hy)) (zero_left : ∀ x hx, p 0 x (zero_mem _) hx) (zero_right : ∀ x hx, p x 0 hx (zero_mem _)) (neg_left : ∀ x y hx hy, p x y hx hy → p (-x) y (neg_mem hx) hy) (neg_right : ∀ x y hx hy, p x y hx hy → p x (-y) hx (neg_mem hy)) (add_left : ∀ x y z hx hy hz, p x z hx hz → p y z hy hz → p (x + y) z (add_mem hx hy) hz) (add_right : ∀ x y z hx hy hz, p x y hx hy → p x z hx hz → p x (y + z) hx (add_mem hy hz)) (mul_left : ∀ x y z hx hy hz, p x z hx hz → p y z hy hz → p (x * y) z (mul_mem hx hy) hz) (mul_right : ∀ x y z hx hy hz, p x y hx hy → p x z hx hz → p x (y * z) hx (mul_mem hy hz)) {x y : R} (hx : x ∈ closure s) (hy : y ∈ closure s) : p x y hx hy
Full source
/-- An induction principle for closure membership, for predicates with two arguments. -/
@[elab_as_elim]
theorem closure_induction₂ {s : Set R} {p : (x y : R) → x ∈ closure sy ∈ closure s → Prop}
    (mem_mem : ∀ (x) (y) (hx : x ∈ s) (hy : y ∈ s), p x y (subset_closure hx) (subset_closure hy))
    (zero_left : ∀ x hx, p 0 x (zero_mem _) hx) (zero_right : ∀ x hx, p x 0 hx (zero_mem _))
    (neg_left : ∀ x y hx hy, p x y hx hy → p (-x) y (neg_mem hx) hy)
    (neg_right : ∀ x y hx hy, p x y hx hy → p x (-y) hx (neg_mem hy))
    (add_left : ∀ x y z hx hy hz, p x z hx hz → p y z hy hz → p (x + y) z (add_mem hx hy) hz)
    (add_right : ∀ x y z hx hy hz, p x y hx hy → p x z hx hz → p x (y + z) hx (add_mem hy hz))
    (mul_left : ∀ x y z hx hy hz, p x z hx hz → p y z hy hz → p (x * y) z (mul_mem hx hy) hz)
    (mul_right : ∀ x y z hx hy hz, p x y hx hy → p x z hx hz → p x (y * z) hx (mul_mem hy hz))
    {x y : R} (hx : x ∈ closure s) (hy : y ∈ closure s) :
    p x y hx hy := by
  induction hy using closure_induction with
  | mem z hz => induction hx using closure_induction with
    | mem _ h => exact mem_mem _ _ h hz
    | zero => exact zero_left _ _
    | mul _ _ _ _ h₁ h₂ => exact mul_left _ _ _ _ _ _ h₁ h₂
    | add _ _ _ _ h₁ h₂ => exact add_left _ _ _ _ _ _ h₁ h₂
    | neg _ _ h => exact neg_left _ _ _ _ h
  | zero => exact zero_right x hx
  | mul _ _ _ _ h₁ h₂ => exact mul_right _ _ _ _ _ _ h₁ h₂
  | add _ _ _ _ h₁ h₂ => exact add_right _ _ _ _ _ _ h₁ h₂
  | neg _ _ h => exact neg_right _ _ _ _ h
Two-Variable Induction Principle for Non-Unital Subring Closure Membership
Informal description
Let $R$ be a non-unital non-associative ring and $s$ a subset of $R$. For a predicate $p$ on pairs of elements of the non-unital subring closure of $s$, if: 1. $p(x,y)$ holds for all $x, y \in s$, 2. $p(0,x)$ and $p(x,0)$ hold for all $x$ in the closure, 3. $p$ is preserved under left and right negation (i.e., $p(x,y)$ implies $p(-x,y)$ and $p(x,-y)$), 4. $p$ is preserved under left and right addition (i.e., $p(x,z)$ and $p(y,z)$ imply $p(x+y,z)$, and $p(x,y)$ and $p(x,z)$ imply $p(x,y+z)$), 5. $p$ is preserved under left and right multiplication (i.e., $p(x,z)$ and $p(y,z)$ imply $p(x \cdot y,z)$, and $p(x,y)$ and $p(x,z)$ imply $p(x,y \cdot z)$), then $p(x,y)$ holds for all $x, y$ in the non-unital subring closure of $s$.
NonUnitalSubring.mem_closure_iff theorem
{s : Set R} {x} : x ∈ closure s ↔ x ∈ AddSubgroup.closure (Subsemigroup.closure s : Set R)
Full source
theorem mem_closure_iff {s : Set R} {x} :
    x ∈ closure sx ∈ closure s ↔ x ∈ AddSubgroup.closure (Subsemigroup.closure s : Set R) :=
  ⟨fun h => by
    induction h using closure_induction with
    | mem _ hx => exact AddSubgroup.subset_closure (Subsemigroup.subset_closure hx)
    | zero => exact zero_mem _
    | add _ _ _ _ hx hy => exact add_mem hx hy
    | neg x _ hx => exact neg_mem hx
    | mul _ _ _hx _hy hx hy =>
      clear _hx _hy
      induction hx, hy using AddSubgroup.closure_induction₂ with
      | mem _ _ hx hy => exact AddSubgroup.subset_closure (mul_mem hx hy)
      | one_left => simpa using zero_mem _
      | one_right => simpa using zero_mem _
      | mul_left _ _ _ _ _ _ h₁ h₂ => simpa [add_mul] using add_mem h₁ h₂
      | mul_right _ _ _ _ _ _ h₁ h₂ => simpa [mul_add] using add_mem h₁ h₂
      | inv_left _ _ _ _ h => simpa [neg_mul] using neg_mem h
      | inv_right _ _ _ _ h => simpa [mul_neg] using neg_mem h,
  fun h => by
    induction h using AddSubgroup.closure_induction with
    | mem _ hx => induction hx using Subsemigroup.closure_induction with
      | mem _ h => exact subset_closure h
      | mul _ _ _ _ h₁ h₂ => exact mul_mem h₁ h₂
    | one => exact zero_mem _
    | mul _ _ _ _ h₁ h₂ => exact add_mem h₁ h₂
    | inv _ _ h => exact neg_mem h⟩
Characterization of Non-Unital Subring Closure Membership via Additive Subgroup and Subsemigroup
Informal description
For any subset $s$ of a non-unital non-associative ring $R$ and any element $x \in R$, $x$ belongs to the non-unital subring closure of $s$ if and only if $x$ belongs to the additive subgroup generated by the subsemigroup generated by $s$.
NonUnitalSubring.closureNonUnitalCommRingOfComm definition
{R : Type u} [NonUnitalRing R] {s : Set R} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : NonUnitalCommRing (closure s)
Full source
/-- If all elements of `s : Set A` commute pairwise, then `closure s` is a commutative ring. -/
def closureNonUnitalCommRingOfComm {R : Type u} [NonUnitalRing R] {s : Set R}
    (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : NonUnitalCommRing (closure s) :=
  { (closure s).toNonUnitalRing with
    mul_comm := fun ⟨x, hx⟩ ⟨y, hy⟩ => by
      ext
      simp only [MulMemClass.mk_mul_mk]
      induction hx, hy using closure_induction₂ with
      | mem_mem x y hx hy => exact hcomm x hx y hy
      | zero_left x _ => exact Commute.zero_left x
      | zero_right x _ => exact Commute.zero_right x
      | mul_left _ _ _ _ _ _ h₁ h₂ => exact Commute.mul_left h₁ h₂
      | mul_right _ _ _ _ _ _ h₁ h₂ => exact Commute.mul_right h₁ h₂
      | add_left _ _ _ _ _ _ h₁ h₂ => exact Commute.add_left h₁ h₂
      | add_right _ _ _ _ _ _ h₁ h₂ => exact Commute.add_right h₁ h₂
      | neg_left _ _ _ _ h => exact Commute.neg_left h
      | neg_right _ _ _ _ h => exact Commute.neg_right h }
Non-unital commutative subring generated by a commuting set
Informal description
Given a non-unital ring $R$ and a subset $s \subseteq R$ where all elements of $s$ commute pairwise (i.e., $a * b = b * a$ for all $a, b \in s$), the non-unital subring generated by $s$ (denoted $\text{closure}(s)$) forms a non-unital commutative ring.
NonUnitalSubring.gi definition
: GaloisInsertion (@closure R _) SetLike.coe
Full source
/-- `closure` forms a Galois insertion with the coercion to set. -/
protected def gi : GaloisInsertion (@closure R _) SetLike.coe where
  choice s _ := closure s
  gc _s _t := closure_le
  le_l_u _s := subset_closure
  choice_eq _s _h := rfl
Galois insertion between subset closure and coercion in non-unital subrings
Informal description
The pair of functions `closure` (which maps a subset of a non-unital ring $R$ to the smallest non-unital subring containing it) and the coercion `coe` (which maps a non-unital subring back to its underlying set) form a Galois insertion. This means that for any subset $s$ of $R$ and any non-unital subring $A$ of $R$, the closure of $s$ is contained in $A$ if and only if $s$ is contained in $A$ (as sets).
NonUnitalSubring.closure_eq theorem
(s : NonUnitalSubring R) : closure (s : Set R) = s
Full source
/-- Closure of a `NonUnitalSubring` `S` equals `S`. -/
@[simp]
theorem closure_eq (s : NonUnitalSubring R) : closure (s : Set R) = s :=
  (NonUnitalSubring.gi R).l_u_eq s
Closure of a Non-Unital Subring Equals Itself
Informal description
For any non-unital subring $S$ of a non-unital non-associative ring $R$, the closure of $S$ (as a subset of $R$) is equal to $S$ itself. That is, $\text{closure}(S) = S$.
NonUnitalSubring.closure_empty theorem
: closure (∅ : Set R) = ⊥
Full source
@[simp]
theorem closure_empty : closure ( : Set R) =  :=
  (NonUnitalSubring.gi R).gc.l_bot
Closure of Empty Set is Trivial Subring
Informal description
The closure of the empty set in a non-unital ring $R$ is the trivial non-unital subring $\{0\}$.
NonUnitalSubring.closure_univ theorem
: closure (Set.univ : Set R) = ⊤
Full source
@[simp]
theorem closure_univ : closure (Set.univ : Set R) =  :=
  @coe_top R _ ▸ closure_eq 
Closure of Universal Set Equals Top Non-Unital Subring
Informal description
The closure of the universal set in a non-unital ring $R$ is equal to the top non-unital subring of $R$, i.e., $\text{closure}(\text{Set.univ}) = \top$.
NonUnitalSubring.closure_union theorem
(s t : Set R) : closure (s ∪ t) = closure s ⊔ closure t
Full source
theorem closure_union (s t : Set R) : closure (s ∪ t) = closureclosure s ⊔ closure t :=
  (NonUnitalSubring.gi R).gc.l_sup
Closure of Union Equals Join of Closures in Non-Unital Subrings
Informal description
For any subsets $s$ and $t$ of a non-unital non-associative ring $R$, the closure of the union $s \cup t$ is equal to the join (supremum) of the closures of $s$ and $t$ in the lattice of non-unital subrings of $R$.
NonUnitalSubring.closure_iUnion theorem
{ι} (s : ι → Set R) : closure (⋃ i, s i) = ⨆ i, closure (s i)
Full source
theorem closure_iUnion {ι} (s : ι → Set R) : closure (⋃ i, s i) = ⨆ i, closure (s i) :=
  (NonUnitalSubring.gi R).gc.l_iSup
Closure of Union Equals Supremum of Closures in Non-Unital Subrings
Informal description
Let $R$ be a non-unital non-associative ring and let $\{s_i\}_{i \in \iota}$ be a family of subsets of $R$. The closure of the union $\bigcup_{i \in \iota} s_i$ is equal to the supremum of the closures of each $s_i$, i.e., \[ \text{closure}\left(\bigcup_{i \in \iota} s_i\right) = \bigsqcup_{i \in \iota} \text{closure}(s_i). \]
NonUnitalSubring.closure_sUnion theorem
(s : Set (Set R)) : closure (⋃₀ s) = ⨆ t ∈ s, closure t
Full source
theorem closure_sUnion (s : Set (Set R)) : closure (⋃₀ s) = ⨆ t ∈ s, closure t :=
  (NonUnitalSubring.gi R).gc.l_sSup
Closure of Union of Subsets in Non-Unital Subrings
Informal description
For any collection $\mathcal{S}$ of subsets of a non-unital non-associative ring $R$, the closure of the union of all sets in $\mathcal{S}$ is equal to the supremum of the closures of each individual set in $\mathcal{S}$. That is, \[ \text{closure}\left(\bigcup_{t \in \mathcal{S}} t\right) = \bigsqcup_{t \in \mathcal{S}} \text{closure}(t). \]
NonUnitalSubring.map_sup theorem
(s t : NonUnitalSubring R) (f : F) : (s ⊔ t).map f = s.map f ⊔ t.map f
Full source
theorem map_sup (s t : NonUnitalSubring R) (f : F) : (s ⊔ t).map f = s.map f ⊔ t.map f :=
  (gc_map_comap f).l_sup
Image of Join of Non-unital Subrings under Homomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be a non-unital ring homomorphism. For any two non-unital subrings $s$ and $t$ of $R$, the image of their join under $f$ equals the join of their images under $f$. That is, \[ f(s \vee t) = f(s) \vee f(t). \]
NonUnitalSubring.map_iSup theorem
{ι : Sort*} (f : F) (s : ι → NonUnitalSubring R) : (iSup s).map f = ⨆ i, (s i).map f
Full source
theorem map_iSup {ι : Sort*} (f : F) (s : ι → NonUnitalSubring R) :
    (iSup s).map f = ⨆ i, (s i).map f :=
  (gc_map_comap f).l_iSup
Image of Supremum of Non-unital Subrings under Ring Homomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be a non-unital ring homomorphism. For any family $\{s_i\}_{i \in \iota}$ of non-unital subrings of $R$, the image of their supremum under $f$ equals the supremum of their images under $f$. That is, \[ f\left(\bigsqcup_{i} s_i\right) = \bigsqcup_{i} f(s_i). \]
NonUnitalSubring.map_inf theorem
(s t : NonUnitalSubring R) (f : F) (hf : Function.Injective f) : (s ⊓ t).map f = s.map f ⊓ t.map f
Full source
theorem map_inf (s t : NonUnitalSubring R) (f : F) (hf : Function.Injective f) :
    (s ⊓ t).map f = s.map f ⊓ t.map f := SetLike.coe_injective (Set.image_inter hf)
Image of Intersection of Non-unital Subrings under Injective Homomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be an injective non-unital ring homomorphism. For any two non-unital subrings $s$ and $t$ of $R$, the image of their intersection under $f$ equals the intersection of their images under $f$. That is, \[ f(s \cap t) = f(s) \cap f(t). \]
NonUnitalSubring.map_iInf theorem
{ι : Sort*} [Nonempty ι] (f : F) (hf : Function.Injective f) (s : ι → NonUnitalSubring R) : (iInf s).map f = ⨅ i, (s i).map f
Full source
theorem map_iInf {ι : Sort*} [Nonempty ι] (f : F) (hf : Function.Injective f)
    (s : ι → NonUnitalSubring R) : (iInf s).map f = ⨅ i, (s i).map f := by
  apply SetLike.coe_injective
  simpa using (Set.injOn_of_injective hf).image_iInter_eq (s := SetLike.coeSetLike.coe ∘ s)
Image of Infimum of Non-unital Subrings under Injective Homomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be an injective non-unital ring homomorphism. For any nonempty family $\{S_i\}_{i \in \iota}$ of non-unital subrings of $R$, the image of their infimum under $f$ equals the infimum of their images under $f$. That is, \[ f\left(\bigsqcap_{i} S_i\right) = \bigsqcap_{i} f(S_i). \]
NonUnitalSubring.comap_inf theorem
(s t : NonUnitalSubring S) (f : F) : (s ⊓ t).comap f = s.comap f ⊓ t.comap f
Full source
theorem comap_inf (s t : NonUnitalSubring S) (f : F) : (s ⊓ t).comap f = s.comap f ⊓ t.comap f :=
  (gc_map_comap f).u_inf
Preimage of Intersection of Non-unital Subrings under Ring Homomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be a non-unital ring homomorphism. For any two non-unital subrings $s$ and $t$ of $S$, the preimage of their intersection under $f$ equals the intersection of their preimages under $f$. That is, \[ f^{-1}(s \cap t) = f^{-1}(s) \cap f^{-1}(t). \]
NonUnitalSubring.comap_iInf theorem
{ι : Sort*} (f : F) (s : ι → NonUnitalSubring S) : (iInf s).comap f = ⨅ i, (s i).comap f
Full source
theorem comap_iInf {ι : Sort*} (f : F) (s : ι → NonUnitalSubring S) :
    (iInf s).comap f = ⨅ i, (s i).comap f :=
  (gc_map_comap f).u_iInf
Preimage of Infimum of Non-Unital Subrings under Ring Homomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be a non-unital ring homomorphism. For any family $\{B_i\}_{i \in \iota}$ of non-unital subrings of $S$, the preimage of their infimum under $f$ equals the infimum of their preimages under $f$. That is, \[ f^{-1}\left(\bigsqcap_{i} B_i\right) = \bigsqcap_{i} f^{-1}(B_i). \]
NonUnitalSubring.map_bot theorem
(f : R →ₙ+* S) : (⊥ : NonUnitalSubring R).map f = ⊥
Full source
@[simp]
theorem map_bot (f : R →ₙ+* S) : ( : NonUnitalSubring R).map f =  :=
  (gc_map_comap f).l_bot
Image of Trivial Subring is Trivial under Non-Unital Ring Homomorphism
Informal description
For any non-unital ring homomorphism $f \colon R \to S$, the image of the trivial non-unital subring $\{0\}$ of $R$ under $f$ is the trivial non-unital subring $\{0\}$ of $S$. That is, $f(\{0\}) = \{0\}$.
NonUnitalSubring.comap_top theorem
(f : R →ₙ+* S) : (⊤ : NonUnitalSubring S).comap f = ⊤
Full source
@[simp]
theorem comap_top (f : R →ₙ+* S) : ( : NonUnitalSubring S).comap f =  :=
  (gc_map_comap f).u_top
Preimage of Top Non-Unital Subring is Top
Informal description
For any non-unital ring homomorphism $f \colon R \to S$, the preimage of the top non-unital subring of $S$ under $f$ is the top non-unital subring of $R$. In other words, $f^{-1}(S) = R$.
NonUnitalSubring.prod definition
(s : NonUnitalSubring R) (t : NonUnitalSubring S) : NonUnitalSubring (R × S)
Full source
/-- Given `NonUnitalSubring`s `s`, `t` of rings `R`, `S` respectively, `s.prod t` is `s ×ˢ t`
as a `NonUnitalSubring` of `R × S`. -/
def prod (s : NonUnitalSubring R) (t : NonUnitalSubring S) : NonUnitalSubring (R × S) :=
  { s.toSubsemigroup.prod t.toSubsemigroup, s.toAddSubgroup.prod t.toAddSubgroup with
    carrier := s ×ˢ t }
Product of non-unital subrings
Informal description
Given non-unital subrings $s$ of $R$ and $t$ of $S$, the product $s \times t$ is a non-unital subring of the product ring $R \times S$. It consists of all pairs $(r, s)$ where $r \in s$ and $s \in t$, and inherits the additive and multiplicative structures from $R \times S$.
NonUnitalSubring.coe_prod theorem
(s : NonUnitalSubring R) (t : NonUnitalSubring S) : (s.prod t : Set (R × S)) = (s : Set R) ×ˢ t
Full source
@[norm_cast]
theorem coe_prod (s : NonUnitalSubring R) (t : NonUnitalSubring S) :
    (s.prod t : Set (R × S)) = (s : Set R) ×ˢ t :=
  rfl
Underlying Set of Product Subring Equals Product of Underlying Sets
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $s$ be a non-unital subring of $R$ and $t$ a non-unital subring of $S$. Then the underlying set of the product subring $s \times t$ is equal to the Cartesian product of the underlying sets $s$ and $t$, i.e., $$(s \times t) = s \times t$$ where on the left side we consider $s \times t$ as a subset of $R \times S$.
NonUnitalSubring.mem_prod theorem
{s : NonUnitalSubring R} {t : NonUnitalSubring S} {p : R × S} : p ∈ s.prod t ↔ p.1 ∈ s ∧ p.2 ∈ t
Full source
theorem mem_prod {s : NonUnitalSubring R} {t : NonUnitalSubring S} {p : R × S} :
    p ∈ s.prod tp ∈ s.prod t ↔ p.1 ∈ s ∧ p.2 ∈ t :=
  Iff.rfl
Membership in Product of Non-unital Subrings
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $s$ and $t$ be non-unital subrings of $R$ and $S$ respectively. For any element $p = (r, s) \in R \times S$, we have $p \in s \times t$ if and only if $r \in s$ and $s \in t$.
NonUnitalSubring.prod_mono theorem
⦃s₁ s₂ : NonUnitalSubring R⦄ (hs : s₁ ≤ s₂) ⦃t₁ t₂ : NonUnitalSubring S⦄ (ht : t₁ ≤ t₂) : s₁.prod t₁ ≤ s₂.prod t₂
Full source
@[mono]
theorem prod_mono ⦃s₁ s₂ : NonUnitalSubring R⦄ (hs : s₁ ≤ s₂) ⦃t₁ t₂ : NonUnitalSubring S⦄
    (ht : t₁ ≤ t₂) : s₁.prod t₁ ≤ s₂.prod t₂ :=
  Set.prod_mono hs ht
Monotonicity of Product of Non-unital Subrings
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $s_1, s_2$ be non-unital subrings of $R$ and $t_1, t_2$ be non-unital subrings of $S$. If $s_1 \leq s_2$ and $t_1 \leq t_2$, then the product subring $s_1 \times t_1$ is contained in the product subring $s_2 \times t_2$.
NonUnitalSubring.prod_mono_right theorem
(s : NonUnitalSubring R) : Monotone fun t : NonUnitalSubring S => s.prod t
Full source
theorem prod_mono_right (s : NonUnitalSubring R) :
    Monotone fun t : NonUnitalSubring S => s.prod t :=
  prod_mono (le_refl s)
Monotonicity of Product with Fixed Left Factor in Non-Unital Subrings
Informal description
For any non-unital subring $s$ of a non-unital ring $R$, the function that maps a non-unital subring $t$ of $S$ to the product subring $s \times t$ is monotone. That is, if $t_1 \leq t_2$ in the lattice of non-unital subrings of $S$, then $s \times t_1 \leq s \times t_2$ in the lattice of non-unital subrings of $R \times S$.
NonUnitalSubring.prod_mono_left theorem
(t : NonUnitalSubring S) : Monotone fun s : NonUnitalSubring R => s.prod t
Full source
theorem prod_mono_left (t : NonUnitalSubring S) : Monotone fun s : NonUnitalSubring R => s.prod t :=
  fun _s₁ _s₂ hs => prod_mono hs (le_refl t)
Monotonicity of Left Factor in Product of Non-Unital Subrings
Informal description
For any non-unital subring $t$ of a non-unital ring $S$, the function that maps a non-unital subring $s$ of $R$ to the product subring $s \times t$ is monotone. That is, if $s_1 \leq s_2$ in the lattice of non-unital subrings of $R$, then $s_1 \times t \leq s_2 \times t$ in the lattice of non-unital subrings of $R \times S$.
NonUnitalSubring.prod_top theorem
(s : NonUnitalSubring R) : s.prod (⊤ : NonUnitalSubring S) = s.comap (NonUnitalRingHom.fst R S)
Full source
theorem prod_top (s : NonUnitalSubring R) :
    s.prod ( : NonUnitalSubring S) = s.comap (NonUnitalRingHom.fst R S) :=
  ext fun x => by simp [mem_prod, MonoidHom.coe_fst]
Product Subring with Top Subring Equals Preimage under First Projection
Informal description
For any non-unital subring $s$ of a non-unital ring $R$, the product subring $s \times S$ (where $S$ is the top non-unital subring of $S$) is equal to the preimage of $s$ under the first projection homomorphism $\mathrm{fst} : R \times S \to R$.
NonUnitalSubring.top_prod theorem
(s : NonUnitalSubring S) : (⊤ : NonUnitalSubring R).prod s = s.comap (NonUnitalRingHom.snd R S)
Full source
theorem top_prod (s : NonUnitalSubring S) :
    ( : NonUnitalSubring R).prod s = s.comap (NonUnitalRingHom.snd R S) :=
  ext fun x => by simp [mem_prod, MonoidHom.coe_snd]
Product of Top Non-Unital Subring with Subring Equals Preimage under Second Projection
Informal description
For any non-unital subring $s$ of a non-unital ring $S$, the product of the top non-unital subring of $R$ with $s$ is equal to the preimage of $s$ under the second projection homomorphism $\text{snd} : R \times S \to S$.
NonUnitalSubring.top_prod_top theorem
: (⊤ : NonUnitalSubring R).prod (⊤ : NonUnitalSubring S) = ⊤
Full source
@[simp]
theorem top_prod_top : ( : NonUnitalSubring R).prod ( : NonUnitalSubring S) =  :=
  (top_prod _).trans <| comap_top _
Product of Top Non-Unital Subrings is Top
Informal description
The product of the top non-unital subring of $R$ and the top non-unital subring of $S$ is equal to the top non-unital subring of $R \times S$.
NonUnitalSubring.prodEquiv definition
(s : NonUnitalSubring R) (t : NonUnitalSubring S) : s.prod t ≃+* s × t
Full source
/-- Product of `NonUnitalSubring`s is isomorphic to their product as rings. -/
def prodEquiv (s : NonUnitalSubring R) (t : NonUnitalSubring S) : s.prod t ≃+* s × t :=
  { Equiv.Set.prod (s : Set R) (t : Set S) with
    map_mul' := fun _ _ => rfl
    map_add' := fun _ _ => rfl }
Isomorphism between product of non-unital subrings and their direct product
Informal description
Given non-unital subrings $s$ of $R$ and $t$ of $S$, the product $s \times t$ is isomorphic as a non-unital ring to the direct product of $s$ and $t$. The isomorphism is given by the canonical set product equivalence between the underlying sets of $s$ and $t$, which preserves both the additive and multiplicative structures.
NonUnitalSubring.mem_iSup_of_directed theorem
{ι} [hι : Nonempty ι] {S : ι → NonUnitalSubring R} (hS : Directed (· ≤ ·) S) {x : R} : (x ∈ ⨆ i, S i) ↔ ∃ i, x ∈ S i
Full source
/-- The underlying set of a non-empty directed Sup of `NonUnitalSubring`s is just a union of the
`NonUnitalSubring`s. Note that this fails without the directedness assumption (the union of two
`NonUnitalSubring`s is typically not a `NonUnitalSubring`) -/
theorem mem_iSup_of_directed {ι} [hι : Nonempty ι] {S : ι → NonUnitalSubring R}
    (hS : Directed (· ≤ ·) S) {x : R} : (x ∈ ⨆ i, S i) ↔ ∃ i, x ∈ S i := by
  refine ⟨?_, fun ⟨i, hi⟩ ↦ le_iSup S i hi⟩
  let U : NonUnitalSubring R :=
    NonUnitalSubring.mk' (⋃ i, (S i : Set R)) (⨆ i, (S i).toSubsemigroup) (⨆ i, (S i).toAddSubgroup)
      (Subsemigroup.coe_iSup_of_directed hS) (AddSubgroup.coe_iSup_of_directed hS)
  suffices ⨆ i, S i ≤ U by simpa [U] using @this x
  exact iSup_le fun i x hx ↦ Set.mem_iUnion.2 ⟨i, hx⟩
Characterization of Membership in Directed Join of Non-Unital Subrings
Informal description
Let $R$ be a non-unital non-associative ring, $\iota$ a non-empty index set, and $\{S_i\}_{i\in\iota}$ a directed family of non-unital subrings of $R$ under inclusion. For any $x \in R$, we have $x \in \bigvee_{i\in\iota} S_i$ if and only if there exists some $i \in \iota$ such that $x \in S_i$. Here, $\bigvee_{i\in\iota} S_i$ denotes the supremum (join) of the family $\{S_i\}$ in the complete lattice of non-unital subrings of $R$.
NonUnitalSubring.coe_iSup_of_directed theorem
{ι} [Nonempty ι] {S : ι → NonUnitalSubring R} (hS : Directed (· ≤ ·) S) : ((⨆ i, S i : NonUnitalSubring R) : Set R) = ⋃ i, S i
Full source
theorem coe_iSup_of_directed {ι} [Nonempty ι] {S : ι → NonUnitalSubring R}
    (hS : Directed (· ≤ ·) S) : ((⨆ i, S i : NonUnitalSubring R) : Set R) = ⋃ i, S i :=
  Set.ext fun x ↦ by simp [mem_iSup_of_directed hS]
Underlying Set of Directed Join of Non-Unital Subrings Equals Union of Subrings
Informal description
Let $R$ be a non-unital non-associative ring, $\iota$ a non-empty index set, and $\{S_i\}_{i\in\iota}$ a directed family of non-unital subrings of $R$ under inclusion. Then the underlying set of the supremum (join) $\bigvee_{i\in\iota} S_i$ in the complete lattice of non-unital subrings of $R$ is equal to the union $\bigcup_{i\in\iota} S_i$ of the underlying sets of the subrings $S_i$.
NonUnitalSubring.mem_sSup_of_directedOn theorem
{S : Set (NonUnitalSubring R)} (Sne : S.Nonempty) (hS : DirectedOn (· ≤ ·) S) {x : R} : x ∈ sSup S ↔ ∃ s ∈ S, x ∈ s
Full source
theorem mem_sSup_of_directedOn {S : Set (NonUnitalSubring R)} (Sne : S.Nonempty)
    (hS : DirectedOn (· ≤ ·) S) {x : R} : x ∈ sSup Sx ∈ sSup S ↔ ∃ s ∈ S, x ∈ s := by
  haveI : Nonempty S := Sne.to_subtype
  simp only [sSup_eq_iSup', mem_iSup_of_directed hS.directed_val, SetCoe.exists, Subtype.coe_mk,
    exists_prop]
Characterization of Membership in Directed Supremum of Non-Unital Subrings
Informal description
Let $R$ be a non-unital non-associative ring and let $S$ be a non-empty collection of non-unital subrings of $R$ that is directed with respect to inclusion. For any element $x \in R$, we have $x \in \bigvee S$ if and only if there exists a subring $s \in S$ such that $x \in s$. Here, $\bigvee S$ denotes the supremum (join) of the collection $S$ in the complete lattice of non-unital subrings of $R$.
NonUnitalSubring.coe_sSup_of_directedOn theorem
{S : Set (NonUnitalSubring R)} (Sne : S.Nonempty) (hS : DirectedOn (· ≤ ·) S) : (↑(sSup S) : Set R) = ⋃ s ∈ S, ↑s
Full source
theorem coe_sSup_of_directedOn {S : Set (NonUnitalSubring R)} (Sne : S.Nonempty)
    (hS : DirectedOn (· ≤ ·) S) : (↑(sSup S) : Set R) = ⋃ s ∈ S, ↑s :=
  Set.ext fun x => by simp [mem_sSup_of_directedOn Sne hS]
Underlying Set of Directed Supremum of Non-Unital Subrings Equals Union
Informal description
Let $R$ be a non-unital non-associative ring and let $S$ be a non-empty collection of non-unital subrings of $R$ that is directed with respect to inclusion. Then the underlying set of the supremum $\bigvee S$ (in the complete lattice of non-unital subrings of $R$) is equal to the union $\bigcup_{s \in S} s$ of the underlying sets of the subrings in $S$.
NonUnitalSubring.mem_map_equiv theorem
{f : R ≃+* S} {K : NonUnitalSubring R} {x : S} : x ∈ K.map (f : R →ₙ+* S) ↔ f.symm x ∈ K
Full source
theorem mem_map_equiv {f : R ≃+* S} {K : NonUnitalSubring R} {x : S} :
    x ∈ K.map (f : R →ₙ+* S)x ∈ K.map (f : R →ₙ+* S) ↔ f.symm x ∈ K :=
  @Set.mem_image_equiv _ _ (K : Set R) f.toEquiv x
Characterization of elements in the image of a subring under a ring isomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be a ring isomorphism. For any non-unital subring $K$ of $R$ and any element $x \in S$, we have $x \in f(K)$ if and only if $f^{-1}(x) \in K$.
NonUnitalSubring.map_equiv_eq_comap_symm theorem
(f : R ≃+* S) (K : NonUnitalSubring R) : K.map (f : R →ₙ+* S) = K.comap f.symm
Full source
theorem map_equiv_eq_comap_symm (f : R ≃+* S) (K : NonUnitalSubring R) :
    K.map (f : R →ₙ+* S) = K.comap f.symm :=
  SetLike.coe_injective (f.toEquiv.image_eq_preimage K)
Image-Preimage Correspondence for Non-unital Subrings under Ring Isomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \simeq S$ be a ring isomorphism. For any non-unital subring $K$ of $R$, the image of $K$ under $f$ is equal to the preimage of $K$ under the inverse isomorphism $f^{-1} \colon S \simeq R$. In other words, $f(K) = f^{-1}(K)$ as non-unital subrings of $S$.
NonUnitalSubring.comap_equiv_eq_map_symm theorem
(f : R ≃+* S) (K : NonUnitalSubring S) : K.comap (f : R →ₙ+* S) = K.map f.symm
Full source
theorem comap_equiv_eq_map_symm (f : R ≃+* S) (K : NonUnitalSubring S) :
    K.comap (f : R →ₙ+* S) = K.map f.symm :=
  (map_equiv_eq_comap_symm f.symm K).symm
Preimage-Image Correspondence for Non-unital Subrings under Ring Isomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \simeq S$ be a ring isomorphism. For any non-unital subring $K$ of $S$, the preimage of $K$ under $f$ is equal to the image of $K$ under the inverse isomorphism $f^{-1} \colon S \simeq R$. In other words, $f^{-1}(K) = f^{-1}(K)$ as non-unital subrings of $R$.
NonUnitalRingHom.rangeRestrict definition
(f : R →ₙ+* S) : R →ₙ+* f.range
Full source
/-- Restriction of a ring homomorphism to its range interpreted as a `NonUnitalSubring`.

This is the bundled version of `Set.rangeFactorization`. -/
def rangeRestrict (f : R →ₙ+* S) : R →ₙ+* f.range :=
  NonUnitalRingHom.codRestrict f f.range fun x => ⟨x, rfl⟩
Range restriction of a non-unital ring homomorphism
Informal description
Given a non-unital ring homomorphism $f \colon R \to S$ between non-unital non-associative rings $R$ and $S$, the range restriction of $f$ is the homomorphism from $R$ to the range of $f$ (viewed as a non-unital subring of $S$) that maps each $x \in R$ to $f(x) \in f.\text{range}$.
NonUnitalRingHom.coe_rangeRestrict theorem
(f : R →ₙ+* S) (x : R) : (f.rangeRestrict x : S) = f x
Full source
@[simp]
theorem coe_rangeRestrict (f : R →ₙ+* S) (x : R) : (f.rangeRestrict x : S) = f x :=
  rfl
Range Restriction Preserves Image: $(f.\text{rangeRestrict}(x) : S) = f(x)$
Informal description
For any non-unital ring homomorphism $f \colon R \to S$ and any element $x \in R$, the image of $x$ under the range-restricted homomorphism $f.\text{rangeRestrict}$ (viewed as an element of $S$) equals $f(x)$.
NonUnitalRingHom.rangeRestrict_surjective theorem
(f : R →ₙ+* S) : Function.Surjective f.rangeRestrict
Full source
theorem rangeRestrict_surjective (f : R →ₙ+* S) : Function.Surjective f.rangeRestrict :=
  fun ⟨_y, hy⟩ =>
  let ⟨x, hx⟩ := mem_range.mp hy
  ⟨x, Subtype.ext hx⟩
Surjectivity of Range-Restricted Non-Unital Ring Homomorphism
Informal description
For any non-unital ring homomorphism $f \colon R \to S$ between non-unital non-associative rings $R$ and $S$, the range restriction $f_{\text{range}} \colon R \to f.\text{range}$ is surjective. That is, for every $y \in f.\text{range}$, there exists an $x \in R$ such that $f_{\text{range}}(x) = y$.
NonUnitalRingHom.range_eq_top theorem
{f : R →ₙ+* S} : f.range = (⊤ : NonUnitalSubring S) ↔ Function.Surjective f
Full source
theorem range_eq_top {f : R →ₙ+* S} :
    f.range = (⊤ : NonUnitalSubring S) ↔ Function.Surjective f :=
  SetLike.ext'_iff.trans <| Iff.trans (by rw [coe_range, coe_top]) Set.range_eq_univ
Range of Non-Unital Ring Homomorphism is Top Subring if and only if Surjective
Informal description
For a non-unital ring homomorphism $f \colon R \to S$, the range of $f$ is equal to the top non-unital subring of $S$ (i.e., $S$ itself) if and only if $f$ is surjective.
NonUnitalRingHom.range_eq_top_of_surjective theorem
(f : R →ₙ+* S) (hf : Function.Surjective f) : f.range = (⊤ : NonUnitalSubring S)
Full source
/-- The range of a surjective ring homomorphism is the whole of the codomain. -/
@[simp]
theorem range_eq_top_of_surjective (f : R →ₙ+* S) (hf : Function.Surjective f) :
    f.range = ( : NonUnitalSubring S) :=
  range_eq_top.2 hf
Range of Surjective Non-Unital Ring Homomorphism is Entire Codomain
Informal description
For a non-unital ring homomorphism $f \colon R \to S$, if $f$ is surjective, then the range of $f$ is equal to the entire codomain $S$ (as a non-unital subring).
NonUnitalRingHom.eqLocus definition
(f g : R →ₙ+* S) : NonUnitalSubring R
Full source
/-- The `NonUnitalSubring` of elements `x : R` such that `f x = g x`, i.e.,
  the equalizer of f and g as a `NonUnitalSubring` of R -/
def eqLocus (f g : R →ₙ+* S) : NonUnitalSubring R :=
  { (f : R →ₙ* S).eqLocus g, (f : R →+ S).eqLocus g with carrier := {x | f x = g x} }
Equalizer subring of non-unital ring homomorphisms
Informal description
Given two non-unital ring homomorphisms $f, g : R \to S$, the equalizer of $f$ and $g$ is the non-unital subring of $R$ consisting of all elements $x \in R$ such that $f(x) = g(x)$. This structure combines the multiplicative and additive equalizers of $f$ and $g$.
NonUnitalRingHom.eqLocus_same theorem
(f : R →ₙ+* S) : f.eqLocus f = ⊤
Full source
@[simp]
theorem eqLocus_same (f : R →ₙ+* S) : f.eqLocus f =  :=
  SetLike.ext fun _ => eq_self_iff_true _
Equalizer subring of a homomorphism with itself is the full ring
Informal description
For any non-unital ring homomorphism $f \colon R \to S$, the equalizer subring of $f$ with itself is the entire ring $R$, i.e., $\text{eqLocus}(f, f) = \top$.
NonUnitalRingHom.eqOn_set_closure theorem
{f g : R →ₙ+* S} {s : Set R} (h : Set.EqOn f g s) : Set.EqOn f g (closure s)
Full source
/-- If two ring homomorphisms are equal on a set, then they are equal on its
`NonUnitalSubring` closure. -/
theorem eqOn_set_closure {f g : R →ₙ+* S} {s : Set R} (h : Set.EqOn f g s) :
    Set.EqOn f g (closure s) :=
  show closure s ≤ f.eqLocus g from closure_le.2 h
Agreement of Non-Unital Ring Homomorphisms on Subring Closure
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f, g \colon R \to S$ be non-unital ring homomorphisms. If $f$ and $g$ agree on a subset $s \subseteq R$, then they also agree on the non-unital subring generated by $s$, i.e., $\forall x \in \text{closure}(s), f(x) = g(x)$.
NonUnitalRingHom.closure_preimage_le theorem
(f : R →ₙ+* S) (s : Set S) : closure (f ⁻¹' s) ≤ (closure s).comap f
Full source
theorem closure_preimage_le (f : R →ₙ+* S) (s : Set S) : closure (f ⁻¹' s) ≤ (closure s).comap f :=
  closure_le.2 fun _x hx => SetLike.mem_coe.2 <| mem_comap.2 <| subset_closure hx
Closure-Preimage Containment for Non-Unital Ring Homomorphisms
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f : R \to S$ be a non-unital ring homomorphism. For any subset $s \subseteq S$, the closure of the preimage $f^{-1}(s)$ in $R$ is contained in the preimage of the closure of $s$ under $f$. In other words, $\text{closure}(f^{-1}(s)) \leq (\text{closure}(s)).\text{comap}\,f$.
NonUnitalRingHom.map_closure theorem
(f : R →ₙ+* S) (s : Set R) : (closure s).map f = closure (f '' s)
Full source
/-- The image under a ring homomorphism of the `NonUnitalSubring` generated by a set equals
the `NonUnitalSubring` generated by the image of the set. -/
theorem map_closure (f : R →ₙ+* S) (s : Set R) : (closure s).map f = closure (f '' s) :=
  Set.image_preimage.l_comm_of_u_comm (gc_map_comap f) (NonUnitalSubring.gi S).gc
    (NonUnitalSubring.gi R).gc fun _ ↦ rfl
Image of Subring Closure Equals Closure of Image under Ring Homomorphism
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be a non-unital ring homomorphism. For any subset $s \subseteq R$, the image under $f$ of the non-unital subring generated by $s$ is equal to the non-unital subring generated by the image $f(s)$. In other words, $f(\text{closure}(s)) = \text{closure}(f(s))$.
NonUnitalSubring.range_subtype theorem
(s : NonUnitalSubring R) : (NonUnitalSubringClass.subtype s).range = s
Full source
@[simp]
theorem range_subtype (s : NonUnitalSubring R) : (NonUnitalSubringClass.subtype s).range = s :=
  SetLike.coe_injective <| (coe_srange _).trans Subtype.range_coe
Range of Inclusion Homomorphism Equals Subring
Informal description
For any non-unital subring $s$ of a non-unital ring $R$, the range of the inclusion homomorphism $\text{subtype} \colon s \to R$ is equal to $s$ itself.
NonUnitalSubring.range_fst theorem
: NonUnitalRingHom.srange (fst R S) = ⊤
Full source
theorem range_fst : NonUnitalRingHom.srange (fst R S) =  :=
  NonUnitalSubsemiring.range_fst
Range of First Projection is Entire Ring
Informal description
The range of the first projection non-unital ring homomorphism from the product ring $R \times S$ to $R$ is the entire ring $R$, i.e., $\text{range}(\text{fst}) = \top$.
NonUnitalSubring.range_snd theorem
: NonUnitalRingHom.srange (snd R S) = ⊤
Full source
theorem range_snd : NonUnitalRingHom.srange (snd R S) =  :=
  NonUnitalSubsemiring.range_snd
Range of Second Projection is Top Non-Unital Subring
Informal description
The range of the second projection non-unital ring homomorphism `snd : R × S → S` is equal to the top non-unital subring of `S`.
RingEquiv.nonUnitalSubringCongr definition
(h : s = t) : s ≃+* t
Full source
/-- Makes the identity isomorphism from a proof two `NonUnitalSubring`s of a multiplicative
    monoid are equal. -/
def nonUnitalSubringCongr (h : s = t) : s ≃+* t :=
  {
    Equiv.setCongr <| congr_arg _ h with
    map_mul' := fun _ _ => rfl
    map_add' := fun _ _ => rfl }
Ring isomorphism between equal non-unital subrings
Informal description
Given two non-unital subrings \( s \) and \( t \) of a non-unital ring \( R \) that are equal (i.e., \( s = t \)), the function constructs a ring isomorphism \( s \simeq+* t \) that preserves both the multiplicative and additive structures. The isomorphism is built using the equivalence of the underlying sets induced by the equality \( s = t \).
RingEquiv.ofLeftInverse' definition
{g : S → R} {f : R →ₙ+* S} (h : Function.LeftInverse g f) : R ≃+* f.range
Full source
/-- Restrict a ring homomorphism with a left inverse to a ring isomorphism to its
`RingHom.range`. -/
def ofLeftInverse' {g : S → R} {f : R →ₙ+* S} (h : Function.LeftInverse g f) : R ≃+* f.range :=
  { f.rangeRestrict with
    toFun := fun x => f.rangeRestrict x
    invFun := fun x => (g ∘ NonUnitalSubringClass.subtype f.range) x
    left_inv := h
    right_inv := fun x =>
      Subtype.ext <|
        let ⟨x', hx'⟩ := NonUnitalRingHom.mem_range.mp x.prop
        show f (g x) = x by rw [← hx', h x'] }
Ring isomorphism from a left inverse to range
Informal description
Given a non-unital ring homomorphism $f \colon R \to S$ between non-unital non-associative rings $R$ and $S$, and a left inverse $g \colon S \to R$ of $f$ (i.e., $g \circ f = \text{id}_R$), the function constructs a ring isomorphism between $R$ and the range of $f$ (viewed as a non-unital subring of $S$). The isomorphism maps each $x \in R$ to $f(x) \in f.\text{range}$, and its inverse maps each $y \in f.\text{range}$ to $g(y) \in R$.
RingEquiv.ofLeftInverse'_apply theorem
{g : S → R} {f : R →ₙ+* S} (h : Function.LeftInverse g f) (x : R) : ↑(ofLeftInverse' h x) = f x
Full source
@[simp]
theorem ofLeftInverse'_apply {g : S → R} {f : R →ₙ+* S} (h : Function.LeftInverse g f) (x : R) :
    ↑(ofLeftInverse' h x) = f x :=
  rfl
Image under Left Inverse-Induced Ring Isomorphism Equals Original Map Value
Informal description
Let $R$ and $S$ be non-unital non-associative rings, $f \colon R \to S$ be a non-unital ring homomorphism, and $g \colon S \to R$ be a left inverse of $f$ (i.e., $g \circ f = \text{id}_R$). For any $x \in R$, the image of $x$ under the ring isomorphism $\text{ofLeftInverse}'\, h \colon R \simeq+* f.\text{range}$ equals $f(x)$ in $S$. In other words, if $\varphi = \text{ofLeftInverse}'\, h$ is the isomorphism constructed from the left inverse $g$, then for any $x \in R$, we have $\varphi(x) = f(x)$ when viewed as elements of $S$.
RingEquiv.ofLeftInverse'_symm_apply theorem
{g : S → R} {f : R →ₙ+* S} (h : Function.LeftInverse g f) (x : f.range) : (ofLeftInverse' h).symm x = g x
Full source
@[simp]
theorem ofLeftInverse'_symm_apply {g : S → R} {f : R →ₙ+* S} (h : Function.LeftInverse g f)
    (x : f.range) : (ofLeftInverse' h).symm x = g x :=
  rfl
Inverse of Ring Isomorphism from Left Inverse to Range
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f \colon R \to S$ be a non-unital ring homomorphism. Given a left inverse $g \colon S \to R$ of $f$ (i.e., $g \circ f = \text{id}_R$), the inverse of the ring isomorphism $\text{ofLeftInverse}'\, h \colon R \simeq f.\text{range}$ satisfies $(\text{ofLeftInverse}'\, h)^{-1}(x) = g(x)$ for all $x$ in the range of $f$.
NonUnitalSubring.closure_preimage_le theorem
(f : F) (s : Set S) : closure ((f : R → S) ⁻¹' s) ≤ (closure s).comap f
Full source
theorem closure_preimage_le (f : F) (s : Set S) :
    closure ((f : R → S) ⁻¹' s) ≤ (closure s).comap f :=
  closure_le.2 fun _x hx => SetLike.mem_coe.2 <| mem_comap.2 <| subset_closure hx
Inclusion of Preimage Closure in Closure Preimage for Non-Unital Subrings
Informal description
Let $R$ and $S$ be non-unital non-associative rings, and let $f : R \to S$ be a non-unital ring homomorphism. For any subset $s \subseteq S$, the non-unital subring generated by the preimage $f^{-1}(s)$ is contained in the preimage of the non-unital subring generated by $s$ under $f$. In other words: $$\text{closure}(f^{-1}(s)) \leq (\text{closure}(s)).\text{comap}(f)$$