Documentation

Mathlib.RingTheory.Subring.Basic

Subrings #

Let R be a ring. This file defines the "bundled" subring type Subring R, a type whose terms correspond to subrings of R. This is the preferred way to talk about subrings in mathlib. Unbundled subrings (s : Set R and IsSubring s) are not in this file, and they will ultimately be deprecated.

We prove that 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 Subring R, sending a subset of R to the subring it generates, and prove that it is a Galois insertion.

Main definitions #

Notation used here:

(R : Type u) [Ring R] (S : Type u) [Ring S] (f g : R →+* S) (A : Subring R) (B : Subring S) (s : Set R)

Implementation notes #

A subring is implemented as a subsemiring which is also an additive subgroup. The initial PR was as a submonoid 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 subring's underlying set.

Tags #

subring, subrings

class SubringClass (S : Type u_1) (R : Type u) [Ring R] [SetLike S R] extends SubsemiringClass , NegMemClass :

    SubringClass S R states that S is a type of subsets s ⊆ R that are both a multiplicative submonoid and an additive subgroup.

    Instances
    theorem coe_int_mem {R : Type u} {S : Type v} [Ring R] [SetLike S R] [hSR : SubringClass S R] (s : S) (n : ) :
    n s
    instance SubringClass.toHasIntCast {R : Type u} {S : Type v} [Ring R] [SetLike S R] [hSR : SubringClass S R] (s : S) :
    IntCast { x // x s }
    Equations
    instance SubringClass.toRing {R : Type u} {S : Type v} [Ring R] [SetLike S R] [hSR : SubringClass S R] (s : S) :
    Ring { x // x s }

    A subring of a ring inherits a ring structure

    Equations
    • One or more equations did not get rendered due to their size.
    instance SubringClass.toCommRing {S : Type v} (s : S) {R : Type u_1} [CommRing R] [SetLike S R] [SubringClass S R] :
    CommRing { x // x s }

    A subring of a CommRing is a CommRing.

    Equations
    • One or more equations did not get rendered due to their size.
    instance SubringClass.toOrderedRing {S : Type v} (s : S) {R : Type u_1} [OrderedRing R] [SetLike S R] [SubringClass S R] :
    OrderedRing { x // x s }

    A subring of an OrderedRing is an OrderedRing.

    Equations
    • One or more equations did not get rendered due to their size.
    instance SubringClass.toOrderedCommRing {S : Type v} (s : S) {R : Type u_1} [OrderedCommRing R] [SetLike S R] [SubringClass S R] :
    OrderedCommRing { x // x s }

    A subring of an OrderedCommRing is an OrderedCommRing.

    Equations
    • One or more equations did not get rendered due to their size.
    instance SubringClass.toLinearOrderedRing {S : Type v} (s : S) {R : Type u_1} [LinearOrderedRing R] [SetLike S R] [SubringClass S R] :
    LinearOrderedRing { x // x s }

    A subring of a LinearOrderedRing is a LinearOrderedRing.

    Equations
    • One or more equations did not get rendered due to their size.
    instance SubringClass.toLinearOrderedCommRing {S : Type v} (s : S) {R : Type u_1} [LinearOrderedCommRing R] [SetLike S R] [SubringClass S R] :

    A subring of a LinearOrderedCommRing is a LinearOrderedCommRing.

    Equations
    • One or more equations did not get rendered due to their size.
    def SubringClass.subtype {R : Type u} {S : Type v} [Ring R] [SetLike S R] [hSR : SubringClass S R] (s : S) :
    { x // x s } →+* R

    The natural ring hom from a subring of ring R to R.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem SubringClass.coeSubtype {R : Type u} {S : Type v} [Ring R] [SetLike S R] [hSR : SubringClass S R] (s : S) :
      ↑(SubringClass.subtype s) = Subtype.val
      @[simp]
      theorem SubringClass.coe_natCast {R : Type u} {S : Type v} [Ring R] [SetLike S R] [hSR : SubringClass S R] (s : S) (n : ) :
      n = n
      @[simp]
      theorem SubringClass.coe_intCast {R : Type u} {S : Type v} [Ring R] [SetLike S R] [hSR : SubringClass S R] (s : S) (n : ) :
      n = n
      structure Subring (R : Type u) [Ring R] extends Subsemiring :
      • carrier : Set R
      • mul_mem' : ∀ {a b : R}, a s.carrierb s.carriera * b s.carrier
      • one_mem' : 1 s.carrier
      • add_mem' : ∀ {a b : R}, a s.carrierb s.carriera + b s.carrier
      • zero_mem' : 0 s.carrier
      • neg_mem' : ∀ {x : R}, x s.carrier-x s.carrier

        G is closed under negation

      Subring R is the type of subrings of R. A subring of R is a subset s that is a multiplicative submonoid and an additive subgroup. Note in particular that it shares the same 0 and 1 as R.

      Instances For
      @[reducible]
      abbrev Subring.toAddSubgroup {R : Type u} [Ring R] (self : Subring R) :

      Reinterpret a Subring as an AddSubgroup.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        instance Subring.instSetLikeSubring {R : Type u} [Ring R] :
        Equations
        • Subring.instSetLikeSubring = { coe := fun s => s.carrier, coe_injective' := (_ : ∀ (p q : Subring R), (fun s => s.carrier) p = (fun s => s.carrier) qp = q) }
        @[simp]
        theorem Subring.mem_toSubsemiring {R : Type u} [Ring R] {s : Subring R} {x : R} :
        x s.toSubsemiring x s
        theorem Subring.mem_carrier {R : Type u} [Ring R] {s : Subring R} {x : R} :
        x s.carrier x s
        @[simp]
        theorem Subring.mem_mk {R : Type u} [Ring R] {S : Subsemiring R} {x : R} (h : ∀ {x : R}, x S.carrier-x S.carrier) :
        x { toSubsemiring := S, neg_mem' := h } x S
        @[simp]
        theorem Subring.coe_set_mk {R : Type u} [Ring R] (S : Subsemiring R) (h : ∀ {x : R}, x S.carrier-x S.carrier) :
        { toSubsemiring := S, neg_mem' := h } = S
        @[simp]
        theorem Subring.mk_le_mk {R : Type u} [Ring R] {S : Subsemiring R} {S' : Subsemiring R} (h₁ : ∀ {x : R}, x S.carrier-x S.carrier) (h₂ : ∀ {x : R}, x S'.carrier-x S'.carrier) :
        { toSubsemiring := S, neg_mem' := h₁ } { toSubsemiring := S', neg_mem' := h₂ } S S'
        theorem Subring.ext {R : Type u} [Ring R] {S : Subring R} {T : Subring R} (h : ∀ (x : R), x S x T) :
        S = T

        Two subrings are equal if they have the same elements.

        def Subring.copy {R : Type u} [Ring R] (S : Subring R) (s : Set R) (hs : s = S) :

        Copy of a subring with a new carrier equal to the old one. Useful to fix definitional equalities.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          @[simp]
          theorem Subring.coe_copy {R : Type u} [Ring R] (S : Subring R) (s : Set R) (hs : s = S) :
          ↑(Subring.copy S s hs) = s
          theorem Subring.copy_eq {R : Type u} [Ring R] (S : Subring R) (s : Set R) (hs : s = S) :
          Subring.copy S s hs = S
          theorem Subring.toSubsemiring_injective {R : Type u} [Ring R] :
          Function.Injective Subring.toSubsemiring
          theorem Subring.toSubsemiring_strictMono {R : Type u} [Ring R] :
          StrictMono Subring.toSubsemiring
          theorem Subring.toSubsemiring_mono {R : Type u} [Ring R] :
          Monotone Subring.toSubsemiring
          theorem Subring.toAddSubgroup_injective {R : Type u} [Ring R] :
          Function.Injective Subring.toAddSubgroup
          theorem Subring.toAddSubgroup_strictMono {R : Type u} [Ring R] :
          StrictMono Subring.toAddSubgroup
          theorem Subring.toAddSubgroup_mono {R : Type u} [Ring R] :
          Monotone Subring.toAddSubgroup
          theorem Subring.toSubmonoid_injective {R : Type u} [Ring R] :
          Function.Injective fun s => s.toSubmonoid
          theorem Subring.toSubmonoid_strictMono {R : Type u} [Ring R] :
          StrictMono fun s => s.toSubmonoid
          theorem Subring.toSubmonoid_mono {R : Type u} [Ring R] :
          Monotone fun s => s.toSubmonoid
          def Subring.mk' {R : Type u} [Ring R] (s : Set R) (sm : Submonoid R) (sa : AddSubgroup R) (hm : sm = s) (ha : sa = s) :

          Construct a Subring R from a set s, a submonoid sm, and an additive subgroup sa such that x ∈ s ↔ x ∈ sm ↔ x ∈ sa.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            @[simp]
            theorem Subring.coe_mk' {R : Type u} [Ring R] {s : Set R} {sm : Submonoid R} (hm : sm = s) {sa : AddSubgroup R} (ha : sa = s) :
            ↑(Subring.mk' s sm sa hm ha) = s
            @[simp]
            theorem Subring.mem_mk' {R : Type u} [Ring R] {s : Set R} {sm : Submonoid R} (hm : sm = s) {sa : AddSubgroup R} (ha : sa = s) {x : R} :
            x Subring.mk' s sm sa hm ha x s
            @[simp]
            theorem Subring.mk'_toSubmonoid {R : Type u} [Ring R] {s : Set R} {sm : Submonoid R} (hm : sm = s) {sa : AddSubgroup R} (ha : sa = s) :
            (Subring.mk' s sm sa hm ha).toSubsemiring.toSubmonoid = sm
            @[simp]
            theorem Subring.mk'_toAddSubgroup {R : Type u} [Ring R] {s : Set R} {sm : Submonoid R} (hm : sm = s) {sa : AddSubgroup R} (ha : sa = s) :
            def Subsemiring.toSubring {R : Type u} [Ring R] (s : Subsemiring R) (hneg : -1 s) :

            A Subsemiring containing -1 is a Subring.

            Equations
            Instances For
              theorem Subring.one_mem {R : Type u} [Ring R] (s : Subring R) :
              1 s

              A subring contains the ring's 1.

              theorem Subring.zero_mem {R : Type u} [Ring R] (s : Subring R) :
              0 s

              A subring contains the ring's 0.

              theorem Subring.mul_mem {R : Type u} [Ring R] (s : Subring R) {x : R} {y : R} :
              x sy sx * y s

              A subring is closed under multiplication.

              theorem Subring.add_mem {R : Type u} [Ring R] (s : Subring R) {x : R} {y : R} :
              x sy sx + y s

              A subring is closed under addition.

              theorem Subring.neg_mem {R : Type u} [Ring R] (s : Subring R) {x : R} :
              x s-x s

              A subring is closed under negation.

              theorem Subring.sub_mem {R : Type u} [Ring R] (s : Subring R) {x : R} {y : R} (hx : x s) (hy : y s) :
              x - y s

              A subring is closed under subtraction

              theorem Subring.list_prod_mem {R : Type u} [Ring R] (s : Subring R) {l : List R} :
              (∀ (x : R), x lx s) → List.prod l s

              Product of a list of elements in a subring is in the subring.

              theorem Subring.list_sum_mem {R : Type u} [Ring R] (s : Subring R) {l : List R} :
              (∀ (x : R), x lx s) → List.sum l s

              Sum of a list of elements in a subring is in the subring.

              theorem Subring.multiset_prod_mem {R : Type u_1} [CommRing R] (s : Subring R) (m : Multiset R) :
              (∀ (a : R), a ma s) → Multiset.prod m s

              Product of a multiset of elements in a subring of a CommRing is in the subring.

              theorem Subring.multiset_sum_mem {R : Type u_1} [Ring R] (s : Subring R) (m : Multiset R) :
              (∀ (a : R), a ma s) → Multiset.sum m s

              Sum of a multiset of elements in a Subring of a Ring is in the Subring.

              theorem Subring.prod_mem {R : Type u_1} [CommRing R] (s : Subring R) {ι : Type u_2} {t : Finset ι} {f : ιR} (h : ∀ (c : ι), c tf c s) :
              (Finset.prod t fun i => f i) s

              Product of elements of a subring of a CommRing indexed by a Finset is in the subring.

              theorem Subring.sum_mem {R : Type u_1} [Ring R] (s : Subring R) {ι : Type u_2} {t : Finset ι} {f : ιR} (h : ∀ (c : ι), c tf c s) :
              (Finset.sum t fun i => f i) s

              Sum of elements in a Subring of a Ring indexed by a Finset is in the Subring.

              instance Subring.toRing {R : Type u} [Ring R] (s : Subring R) :
              Ring { x // x s }

              A subring of a ring inherits a ring structure

              Equations
              theorem Subring.zsmul_mem {R : Type u} [Ring R] (s : Subring R) {x : R} (hx : x s) (n : ) :
              n x s
              theorem Subring.pow_mem {R : Type u} [Ring R] (s : Subring R) {x : R} (hx : x s) (n : ) :
              x ^ n s
              @[simp]
              theorem Subring.coe_add {R : Type u} [Ring R] (s : Subring R) (x : { x // x s }) (y : { x // x s }) :
              ↑(x + y) = x + y
              @[simp]
              theorem Subring.coe_neg {R : Type u} [Ring R] (s : Subring R) (x : { x // x s }) :
              ↑(-x) = -x
              @[simp]
              theorem Subring.coe_mul {R : Type u} [Ring R] (s : Subring R) (x : { x // x s }) (y : { x // x s }) :
              ↑(x * y) = x * y
              @[simp]
              theorem Subring.coe_zero {R : Type u} [Ring R] (s : Subring R) :
              0 = 0
              @[simp]
              theorem Subring.coe_one {R : Type u} [Ring R] (s : Subring R) :
              1 = 1
              @[simp]
              theorem Subring.coe_pow {R : Type u} [Ring R] (s : Subring R) (x : { x // x s }) (n : ) :
              ↑(x ^ n) = x ^ n
              theorem Subring.coe_eq_zero_iff {R : Type u} [Ring R] (s : Subring R) {x : { x // x s }} :
              x = 0 x = 0
              instance Subring.toCommRing {R : Type u_1} [CommRing R] (s : Subring R) :
              CommRing { x // x s }

              A subring of a CommRing is a CommRing.

              Equations

              A subring of a domain is a domain.

              Equations
              • One or more equations did not get rendered due to their size.
              instance Subring.toOrderedRing {R : Type u_1} [OrderedRing R] (s : Subring R) :
              OrderedRing { x // x s }

              A subring of an OrderedRing is an OrderedRing.

              Equations
              def Subring.subtype {R : Type u} [Ring R] (s : Subring R) :
              { x // x s } →+* R

              The natural ring hom from a subring of ring R to R.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                @[simp]
                theorem Subring.coeSubtype {R : Type u} [Ring R] (s : Subring R) :
                ↑(Subring.subtype s) = Subtype.val
                theorem Subring.coe_natCast {R : Type u} [Ring R] (s : Subring R) (n : ) :
                n = n
                theorem Subring.coe_intCast {R : Type u} [Ring R] (s : Subring R) (n : ) :
                n = n

                Partial order #

                @[simp]
                theorem Subring.coe_toSubsemiring {R : Type u} [Ring R] (s : Subring R) :
                s.toSubsemiring = s
                @[simp]
                theorem Subring.mem_toSubmonoid {R : Type u} [Ring R] {s : Subring R} {x : R} :
                x s.toSubmonoid x s
                @[simp]
                theorem Subring.coe_toSubmonoid {R : Type u} [Ring R] (s : Subring R) :
                s.toSubmonoid = s
                @[simp]
                theorem Subring.mem_toAddSubgroup {R : Type u} [Ring R] {s : Subring R} {x : R} :
                @[simp]
                theorem Subring.coe_toAddSubgroup {R : Type u} [Ring R] (s : Subring R) :

                top #

                instance Subring.instTopSubring {R : Type u} [Ring R] :

                The subring R of the ring R.

                Equations
                • One or more equations did not get rendered due to their size.
                @[simp]
                theorem Subring.mem_top {R : Type u} [Ring R] (x : R) :
                @[simp]
                theorem Subring.coe_top {R : Type u} [Ring R] :
                = Set.univ
                @[simp]
                theorem Subring.topEquiv_apply {R : Type u} [Ring R] (r : { x // x }) :
                Subring.topEquiv r = r
                @[simp]
                theorem Subring.topEquiv_symm_apply_coe {R : Type u} [Ring R] (r : R) :
                ↑(↑(RingEquiv.symm Subring.topEquiv) r) = r
                def Subring.topEquiv {R : Type u} [Ring R] :
                { x // x } ≃+* R

                The ring equiv between the top element of Subring R and R.

                Equations
                • Subring.topEquiv = Subsemiring.topEquiv
                Instances For

                  comap #

                  def Subring.comap {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (s : Subring S) :

                  The preimage of a subring along a ring homomorphism is a subring.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem Subring.coe_comap {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring S) (f : R →+* S) :
                    ↑(Subring.comap f s) = f ⁻¹' s
                    @[simp]
                    theorem Subring.mem_comap {R : Type u} {S : Type v} [Ring R] [Ring S] {s : Subring S} {f : R →+* S} {x : R} :
                    x Subring.comap f s f x s
                    theorem Subring.comap_comap {R : Type u} {S : Type v} {T : Type w} [Ring R] [Ring S] [Ring T] (s : Subring T) (g : S →+* T) (f : R →+* S) :

                    map #

                    def Subring.map {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (s : Subring R) :

                    The image of a subring along a ring homomorphism is a subring.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[simp]
                      theorem Subring.coe_map {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (s : Subring R) :
                      ↑(Subring.map f s) = f '' s
                      @[simp]
                      theorem Subring.mem_map {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} {s : Subring R} {y : S} :
                      y Subring.map f s x, x s f x = y
                      @[simp]
                      theorem Subring.map_id {R : Type u} [Ring R] (s : Subring R) :
                      theorem Subring.map_map {R : Type u} {S : Type v} {T : Type w} [Ring R] [Ring S] [Ring T] (s : Subring R) (g : S →+* T) (f : R →+* S) :
                      theorem Subring.map_le_iff_le_comap {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} {s : Subring R} {t : Subring S} :
                      theorem Subring.gc_map_comap {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) :
                      noncomputable def Subring.equivMapOfInjective {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) (f : R →+* S) (hf : Function.Injective f) :
                      { x // x s } ≃+* { x // x Subring.map f s }

                      A subring is isomorphic to its image under an injective function

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        @[simp]
                        theorem Subring.coe_equivMapOfInjective_apply {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) (f : R →+* S) (hf : Function.Injective f) (x : { x // x s }) :
                        ↑(↑(Subring.equivMapOfInjective s f hf) x) = f x

                        range #

                        def RingHom.range {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) :

                        The range of a ring homomorphism, as a subring of the target. See Note [range copy pattern].

                        Equations
                        Instances For
                        @[simp]
                        theorem RingHom.coe_range {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) :
                        @[simp]
                        theorem RingHom.mem_range {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} {y : S} :
                        y RingHom.range f x, f x = y
                        theorem RingHom.range_eq_map {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) :
                        theorem RingHom.mem_range_self {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (x : R) :
                        theorem RingHom.map_range {R : Type u} {S : Type v} {T : Type w} [Ring R] [Ring S] [Ring T] (g : S →+* T) (f : R →+* S) :
                        instance RingHom.fintypeRange {R : Type u} {S : Type v} [Ring R] [Ring S] [Fintype R] [DecidableEq S] (f : R →+* S) :

                        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.

                        Equations

                        bot #

                        instance Subring.instBotSubring {R : Type u} [Ring R] :
                        Equations
                        Equations
                        • Subring.instInhabitedSubring = { default := }
                        theorem Subring.coe_bot {R : Type u} [Ring R] :
                        = Set.range Int.cast
                        theorem Subring.mem_bot {R : Type u} [Ring R] {x : R} :
                        x n, n = x

                        inf #

                        instance Subring.instInfSubring {R : Type u} [Ring R] :

                        The inf of two subrings is their intersection.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        @[simp]
                        theorem Subring.coe_inf {R : Type u} [Ring R] (p : Subring R) (p' : Subring R) :
                        ↑(p p') = p p'
                        @[simp]
                        theorem Subring.mem_inf {R : Type u} [Ring R] {p : Subring R} {p' : Subring R} {x : R} :
                        x p p' x p x p'
                        Equations
                        • One or more equations did not get rendered due to their size.
                        @[simp]
                        theorem Subring.coe_sInf {R : Type u} [Ring R] (S : Set (Subring R)) :
                        ↑(sInf S) = ⋂ (s : Subring R) (_ : s S), s
                        theorem Subring.mem_sInf {R : Type u} [Ring R] {S : Set (Subring R)} {x : R} :
                        x sInf S ∀ (p : Subring R), p Sx p
                        @[simp]
                        theorem Subring.coe_iInf {R : Type u} [Ring R] {ι : Sort u_1} {S : ιSubring R} :
                        ↑(⨅ (i : ι), S i) = ⋂ (i : ι), ↑(S i)
                        theorem Subring.mem_iInf {R : Type u} [Ring R] {ι : Sort u_1} {S : ιSubring R} {x : R} :
                        x ⨅ (i : ι), S i ∀ (i : ι), x S i
                        @[simp]
                        theorem Subring.sInf_toSubmonoid {R : Type u} [Ring R] (s : Set (Subring R)) :
                        (sInf s).toSubmonoid = ⨅ (t : Subring R) (_ : t s), t.toSubmonoid
                        @[simp]
                        theorem Subring.sInf_toAddSubgroup {R : Type u} [Ring R] (s : Set (Subring R)) :

                        Subrings of a ring form a complete lattice.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        theorem Subring.eq_top_iff' {R : Type u} [Ring R] (A : Subring R) :
                        A = ∀ (x : R), x A

                        Center of a ring #

                        @[simp]
                        theorem Subring.center_toSubsemiring (R : Type u) [Ring R] :
                        (Subring.center R).toSubsemiring = Subsemiring.center R
                        theorem Subring.mem_center_iff {R : Type u} [Ring R] {z : R} :
                        z Subring.center R ∀ (g : R), g * z = z * g
                        Equations
                        @[simp]

                        The center is commutative.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Equations
                        • One or more equations did not get rendered due to their size.
                        @[simp]
                        theorem Subring.center.coe_inv {K : Type u} [DivisionRing K] (a : { x // x Subring.center K }) :
                        a⁻¹ = (a)⁻¹
                        @[simp]
                        theorem Subring.center.coe_div {K : Type u} [DivisionRing K] (a : { x // x Subring.center K }) (b : { x // x Subring.center K }) :
                        ↑(a / b) = a / b
                        def Subring.centralizer {R : Type u} [Ring R] (s : Set R) :

                        The centralizer of a set inside a ring as a Subring.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          @[simp]
                          theorem Subring.centralizer_toSubmonoid {R : Type u} [Ring R] (s : Set R) :
                          (Subring.centralizer s).toSubsemiring.toSubmonoid = Submonoid.centralizer s
                          theorem Subring.mem_centralizer_iff {R : Type u} [Ring R] {s : Set R} {z : R} :
                          z Subring.centralizer s ∀ (g : R), g sg * z = z * g
                          theorem Subring.centralizer_le {R : Type u} [Ring R] (s : Set R) (t : Set R) (h : s t) :

                          subring closure of a subset #

                          def Subring.closure {R : Type u} [Ring R] (s : Set R) :

                          The Subring generated by a set.

                          Equations
                          Instances For
                            theorem Subring.mem_closure {R : Type u} [Ring R] {x : R} {s : Set R} :
                            x Subring.closure s ∀ (S : Subring R), s Sx S
                            @[simp]
                            theorem Subring.subset_closure {R : Type u} [Ring R] {s : Set R} :

                            The subring generated by a set includes the set.

                            theorem Subring.not_mem_of_not_mem_closure {R : Type u} [Ring R] {s : Set R} {P : R} (hP : ¬P Subring.closure s) :
                            ¬P s
                            @[simp]
                            theorem Subring.closure_le {R : Type u} [Ring R] {s : Set R} {t : Subring R} :

                            A subring t includes closure s if and only if it includes s.

                            theorem Subring.closure_mono {R : Type u} [Ring R] ⦃s : Set R ⦃t : Set R (h : s t) :

                            Subring closure of a set is monotone in its argument: if s ⊆ t, then closure s ≤ closure t.

                            theorem Subring.closure_eq_of_le {R : Type u} [Ring R] {s : Set R} {t : Subring R} (h₁ : s t) (h₂ : t Subring.closure s) :
                            theorem Subring.closure_induction {R : Type u} [Ring R] {s : Set R} {p : RProp} {x : R} (h : x Subring.closure s) (Hs : (x : R) → x sp x) (H0 : p 0) (H1 : p 1) (Hadd : (x y : R) → p xp yp (x + y)) (Hneg : (x : R) → p xp (-x)) (Hmul : (x y : R) → p xp yp (x * y)) :
                            p x

                            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.

                            theorem Subring.closure_induction' {R : Type u} [Ring R] {s : Set R} {p : (x : R) → x Subring.closure sProp} (Hs : (x : R) → (h : x s) → p x (_ : x ↑(Subring.closure s))) (H0 : p 0 (_ : 0 Subring.closure s)) (H1 : p 1 (_ : 1 Subring.closure s)) (Hadd : (x : R) → (hx : x Subring.closure s) → (y : R) → (hy : y Subring.closure s) → p x hxp y hyp (x + y) (_ : x + y Subring.closure s)) (Hneg : (x : R) → (hx : x Subring.closure s) → p x hxp (-x) (_ : -x Subring.closure s)) (Hmul : (x : R) → (hx : x Subring.closure s) → (y : R) → (hy : y Subring.closure s) → p x hxp y hyp (x * y) (_ : x * y Subring.closure s)) {a : R} (ha : a Subring.closure s) :
                            p a ha
                            theorem Subring.closure_induction₂ {R : Type u} [Ring R] {s : Set R} {p : RRProp} {a : R} {b : R} (ha : a Subring.closure s) (hb : b Subring.closure s) (Hs : (x : R) → x s(y : R) → y sp x y) (H0_left : (x : R) → p 0 x) (H0_right : (x : R) → p x 0) (H1_left : (x : R) → p 1 x) (H1_right : (x : R) → p x 1) (Hneg_left : (x y : R) → p x yp (-x) y) (Hneg_right : (x y : R) → p x yp x (-y)) (Hadd_left : (x₁ x₂ y : R) → p x₁ yp x₂ yp (x₁ + x₂) y) (Hadd_right : (x y₁ y₂ : R) → p x y₁p x y₂p x (y₁ + y₂)) (Hmul_left : (x₁ x₂ y : R) → p x₁ yp x₂ yp (x₁ * x₂) y) (Hmul_right : (x y₁ y₂ : R) → p x y₁p x y₂p x (y₁ * y₂)) :
                            p a b

                            An induction principle for closure membership, for predicates with two arguments.

                            def Subring.closureCommRingOfComm {R : Type u} [Ring R] {s : Set R} (hcomm : ∀ (a : R), a s∀ (b : R), b sa * b = b * a) :

                            If all elements of s : Set A commute pairwise, then closure s is a commutative ring.

                            Equations
                            Instances For
                              theorem Subring.exists_list_of_mem_closure {R : Type u} [Ring R] {s : Set R} {x : R} (h : x Subring.closure s) :
                              L, (∀ (t : List R), t L∀ (y : R), y ty s y = -1) List.sum (List.map List.prod L) = x
                              def Subring.gi (R : Type u) [Ring R] :
                              GaloisInsertion Subring.closure SetLike.coe

                              closure forms a Galois insertion with the coercion to set.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                theorem Subring.closure_eq {R : Type u} [Ring R] (s : Subring R) :

                                Closure of a subring S equals S.

                                @[simp]
                                theorem Subring.closure_univ {R : Type u} [Ring R] :
                                theorem Subring.closure_iUnion {R : Type u} [Ring R] {ι : Sort u_1} (s : ιSet R) :
                                Subring.closure (⋃ (i : ι), s i) = ⨆ (i : ι), Subring.closure (s i)
                                theorem Subring.closure_sUnion {R : Type u} [Ring R] (s : Set (Set R)) :
                                Subring.closure (⋃₀ s) = ⨆ (t : Set R) (_ : t s), Subring.closure t
                                theorem Subring.map_sup {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) (t : Subring R) (f : R →+* S) :
                                theorem Subring.map_iSup {R : Type u} {S : Type v} [Ring R] [Ring S] {ι : Sort u_1} (f : R →+* S) (s : ιSubring R) :
                                Subring.map f (iSup s) = ⨆ (i : ι), Subring.map f (s i)
                                theorem Subring.comap_inf {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring S) (t : Subring S) (f : R →+* S) :
                                theorem Subring.comap_iInf {R : Type u} {S : Type v} [Ring R] [Ring S] {ι : Sort u_1} (f : R →+* S) (s : ιSubring S) :
                                Subring.comap f (iInf s) = ⨅ (i : ι), Subring.comap f (s i)
                                @[simp]
                                theorem Subring.map_bot {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) :
                                @[simp]
                                theorem Subring.comap_top {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) :
                                def Subring.prod {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) (t : Subring S) :
                                Subring (R × S)

                                Given Subrings s, t of rings R, S respectively, s.prod t is s ×̂ t as a subring of R × S.

                                Equations
                                • One or more equations did not get rendered due to their size.
                                Instances For
                                  theorem Subring.coe_prod {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) (t : Subring S) :
                                  ↑(Subring.prod s t) = s ×ˢ t
                                  theorem Subring.mem_prod {R : Type u} {S : Type v} [Ring R] [Ring S] {s : Subring R} {t : Subring S} {p : R × S} :
                                  p Subring.prod s t p.fst s p.snd t
                                  theorem Subring.prod_mono {R : Type u} {S : Type v} [Ring R] [Ring S] ⦃s₁ : Subring R ⦃s₂ : Subring R (hs : s₁ s₂) ⦃t₁ : Subring S ⦃t₂ : Subring S (ht : t₁ t₂) :
                                  Subring.prod s₁ t₁ Subring.prod s₂ t₂
                                  theorem Subring.prod_mono_right {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) :
                                  Monotone fun t => Subring.prod s t
                                  theorem Subring.prod_mono_left {R : Type u} {S : Type v} [Ring R] [Ring S] (t : Subring S) :
                                  Monotone fun s => Subring.prod s t
                                  theorem Subring.prod_top {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) :
                                  theorem Subring.top_prod {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring S) :
                                  @[simp]
                                  theorem Subring.top_prod_top {R : Type u} {S : Type v} [Ring R] [Ring S] :
                                  def Subring.prodEquiv {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) (t : Subring S) :
                                  { x // x Subring.prod s t } ≃+* { x // x s } × { x // x t }

                                  Product of subrings is isomorphic to their product as rings.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    theorem Subring.mem_iSup_of_directed {R : Type u} [Ring R] {ι : Sort u_1} [hι : Nonempty ι] {S : ιSubring R} (hS : Directed (fun x x_1 => x x_1) S) {x : R} :
                                    x ⨆ (i : ι), S i i, x S i

                                    The underlying set of a non-empty directed sSup of subrings is just a union of the subrings. Note that this fails without the directedness assumption (the union of two subrings is typically not a subring)

                                    theorem Subring.coe_iSup_of_directed {R : Type u} [Ring R] {ι : Sort u_1} [hι : Nonempty ι] {S : ιSubring R} (hS : Directed (fun x x_1 => x x_1) S) :
                                    ↑(⨆ (i : ι), S i) = ⋃ (i : ι), ↑(S i)
                                    theorem Subring.mem_sSup_of_directedOn {R : Type u} [Ring R] {S : Set (Subring R)} (Sne : Set.Nonempty S) (hS : DirectedOn (fun x x_1 => x x_1) S) {x : R} :
                                    x sSup S s, s S x s
                                    theorem Subring.coe_sSup_of_directedOn {R : Type u} [Ring R] {S : Set (Subring R)} (Sne : Set.Nonempty S) (hS : DirectedOn (fun x x_1 => x x_1) S) :
                                    ↑(sSup S) = ⋃ (s : Subring R) (_ : s S), s
                                    theorem Subring.mem_map_equiv {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R ≃+* S} {K : Subring R} {x : S} :
                                    x Subring.map (f) K ↑(RingEquiv.symm f) x K
                                    theorem Subring.map_equiv_eq_comap_symm {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R ≃+* S) (K : Subring R) :
                                    theorem Subring.comap_equiv_eq_map_symm {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R ≃+* S) (K : Subring S) :
                                    def RingHom.rangeRestrict {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) :
                                    R →+* { x // x RingHom.range f }

                                    Restriction of a ring homomorphism to its range interpreted as a subsemiring.

                                    This is the bundled version of Set.rangeFactorization.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem RingHom.coe_rangeRestrict {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (x : R) :
                                      ↑(↑(RingHom.rangeRestrict f) x) = f x
                                      @[simp]
                                      theorem RingHom.range_top_of_surjective {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (hf : Function.Surjective f) :

                                      The range of a surjective ring homomorphism is the whole of the codomain.

                                      def RingHom.eqLocus {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (g : R →+* S) :

                                      The subring of elements x : R such that f x = g x, i.e., the equalizer of f and g as a subring of R

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For
                                        @[simp]
                                        theorem RingHom.eqLocus_same {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) :
                                        theorem RingHom.eqOn_set_closure {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} {g : R →+* S} {s : Set R} (h : Set.EqOn (f) (g) s) :
                                        Set.EqOn f g ↑(Subring.closure s)

                                        If two ring homomorphisms are equal on a set, then they are equal on its subring closure.

                                        theorem RingHom.eq_of_eqOn_set_top {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} {g : R →+* S} (h : Set.EqOn f g ) :
                                        f = g
                                        theorem RingHom.eq_of_eqOn_set_dense {R : Type u} {S : Type v} [Ring R] [Ring S] {s : Set R} (hs : Subring.closure s = ) {f : R →+* S} {g : R →+* S} (h : Set.EqOn (f) (g) s) :
                                        f = g
                                        theorem RingHom.closure_preimage_le {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (s : Set S) :
                                        theorem RingHom.map_closure {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (s : Set R) :

                                        The image under a ring homomorphism of the subring generated by a set equals the subring generated by the image of the set.

                                        def Subring.inclusion {R : Type u} [Ring R] {S : Subring R} {T : Subring R} (h : S T) :
                                        { x // x S } →+* { x // x T }

                                        The ring homomorphism associated to an inclusion of subrings.

                                        Equations
                                        Instances For
                                          @[simp]
                                          theorem Subring.range_fst {R : Type u} {S : Type v} [Ring R] [Ring S] :
                                          theorem Subring.range_snd {R : Type u} {S : Type v} [Ring R] [Ring S] :
                                          @[simp]
                                          theorem Subring.prod_bot_sup_bot_prod {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) (t : Subring S) :
                                          def RingEquiv.subringCongr {R : Type u} [Ring R] {s : Subring R} {t : Subring R} (h : s = t) :
                                          { x // x s } ≃+* { x // x t }

                                          Makes the identity isomorphism from a proof two subrings of a multiplicative monoid are equal.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            def RingEquiv.ofLeftInverse {R : Type u} {S : Type v} [Ring R] [Ring S] {g : SR} {f : R →+* S} (h : Function.LeftInverse g f) :
                                            R ≃+* { x // x RingHom.range f }

                                            Restrict a ring homomorphism with a left inverse to a ring isomorphism to its RingHom.range.

                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For
                                              @[simp]
                                              theorem RingEquiv.ofLeftInverse_apply {R : Type u} {S : Type v} [Ring R] [Ring S] {g : SR} {f : R →+* S} (h : Function.LeftInverse g f) (x : R) :
                                              ↑(↑(RingEquiv.ofLeftInverse h) x) = f x
                                              @[simp]
                                              theorem RingEquiv.ofLeftInverse_symm_apply {R : Type u} {S : Type v} [Ring R] [Ring S] {g : SR} {f : R →+* S} (h : Function.LeftInverse g f) (x : { x // x RingHom.range f }) :
                                              @[simp]
                                              theorem RingEquiv.subringMap_symm_apply_coe {R : Type u} {S : Type v} [Ring R] [Ring S] {s : Subring R} (e : R ≃+* S) :
                                              ∀ (a : ↑(↑(RingEquiv.toAddEquiv e) '' ↑(Subsemiring.toAddSubmonoid s.toSubsemiring))), ↑(↑(RingEquiv.symm (RingEquiv.subringMap e)) a) = ↑(AddEquiv.symm e) a
                                              @[simp]
                                              theorem RingEquiv.subringMap_apply_coe {R : Type u} {S : Type v} [Ring R] [Ring S] {s : Subring R} (e : R ≃+* S) :
                                              ∀ (a : ↑(Subsemiring.toAddSubmonoid s.toSubsemiring)), ↑(↑(RingEquiv.subringMap e) a) = e a
                                              def RingEquiv.subringMap {R : Type u} {S : Type v} [Ring R] [Ring S] {s : Subring R} (e : R ≃+* S) :
                                              { x // x s } ≃+* { x // x Subring.map (RingEquiv.toRingHom e) s }

                                              Given an equivalence e : R ≃+* S of rings and a subring s of R, subringMap e s is the induced equivalence between s and s.map e

                                              Equations
                                              Instances For
                                                theorem Subring.InClosure.recOn {R : Type u} [Ring R] {s : Set R} {C : RProp} {x : R} (hx : x Subring.closure s) (h1 : C 1) (hneg1 : C (-1)) (hs : (z : R) → z s(n : R) → C nC (z * n)) (ha : {x y : R} → C xC yC (x + y)) :
                                                C x
                                                theorem Subring.closure_preimage_le {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (s : Set S) :
                                                theorem AddSubgroup.int_mul_mem {R : Type u} [Ring R] {G : AddSubgroup R} (k : ) {g : R} (h : g G) :
                                                k * g G

                                                Actions by Subrings #

                                                These are just copies of the definitions about Subsemiring starting from Subsemiring.MulAction.

                                                When R is commutative, Algebra.ofSubring provides a stronger result than those found in this file, which uses the same scalar action.

                                                instance Subring.instSMulSubtypeMemSubringInstMembershipInstSetLikeSubring {R : Type u} [Ring R] {α : Type u_1} [SMul R α] (S : Subring R) :
                                                SMul { x // x S } α

                                                The action by a subring is the action by the underlying ring.

                                                Equations
                                                theorem Subring.smul_def {R : Type u} [Ring R] {α : Type u_1} [SMul R α] {S : Subring R} (g : { x // x S }) (m : α) :
                                                g m = g m
                                                instance Subring.smulCommClass_left {R : Type u} [Ring R] {α : Type u_1} {β : Type u_2} [SMul R β] [SMul α β] [SMulCommClass R α β] (S : Subring R) :
                                                SMulCommClass { x // x S } α β
                                                Equations
                                                instance Subring.smulCommClass_right {R : Type u} [Ring R] {α : Type u_1} {β : Type u_2} [SMul α β] [SMul R β] [SMulCommClass α R β] (S : Subring R) :
                                                SMulCommClass α { x // x S } β
                                                Equations

                                                Note that this provides IsScalarTower S R R which is needed by smul_mul_assoc.

                                                Equations
                                                • One or more equations did not get rendered due to their size.
                                                Equations
                                                • One or more equations did not get rendered due to their size.

                                                The action by a subring is the action by the underlying ring.

                                                Equations
                                                • One or more equations did not get rendered due to their size.

                                                The action by a subring is the action by the underlying ring.

                                                Equations
                                                • One or more equations did not get rendered due to their size.

                                                The action by a subring is the action by the underlying ring.

                                                Equations
                                                • One or more equations did not get rendered due to their size.

                                                The action by a subsemiring is the action by the underlying ring.

                                                Equations
                                                • One or more equations did not get rendered due to their size.

                                                The center of a semiring acts commutatively on that semiring.

                                                Equations

                                                The center of a semiring acts commutatively on that semiring.

                                                Equations

                                                The subgroup of positive units of a linear ordered semiring.

                                                Equations
                                                • One or more equations did not get rendered due to their size.
                                                Instances For
                                                  @[simp]
                                                  theorem Units.mem_posSubgroup {R : Type u_1} [LinearOrderedSemiring R] (u : Rˣ) :