Documentation

Mathlib.Algebra.Ring.Subring.Basic

Subrings #

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

theorem Subring.toSubmonoid_mono {R : Type u} [Ring R] :
Monotone fun (s : Subring R) => s.toSubmonoid
theorem Subring.list_prod_mem {R : Type u} [Ring R] (s : Subring R) {l : List R} :
(∀ xl, x s)l.prod 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} :
(∀ xl, x s)l.sum 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) :
(∀ am, a s)m.prod 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) :
(∀ am, a s)m.sum 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 : ct, f c s) :
it, 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 : ct, f c s) :
it, f i s

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

top #

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

The subring R of the ring R.

Equations
  • Subring.instTop = { top := let __src := ; let __src_1 := ; { toSubmonoid := __src, add_mem' := , zero_mem' := , neg_mem' := } }
@[simp]
theorem Subring.mem_top {R : Type u} [Ring R] (x : R) :
@[simp]
theorem Subring.coe_top {R : Type u} [Ring R] :
def Subring.topEquiv {R : Type u} [Ring R] :
≃+* R

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

Equations
Instances For
    @[simp]
    theorem Subring.topEquiv_apply {R : Type u} [Ring R] (r : ) :
    topEquiv r = r
    @[simp]
    theorem Subring.topEquiv_symm_apply_coe {R : Type u} [Ring R] (r : R) :
    (topEquiv.symm r) = r

    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
    • Subring.comap f s = { carrier := f ⁻¹' s.carrier, mul_mem' := , one_mem' := , add_mem' := , zero_mem' := , neg_mem' := }
    Instances For
      @[simp]
      theorem Subring.coe_comap {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring S) (f : R →+* S) :
      (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 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) :
      comap f (comap g s) = comap (g.comp f) 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
      • Subring.map f s = { carrier := f '' s.carrier, mul_mem' := , one_mem' := , add_mem' := , zero_mem' := , neg_mem' := }
      Instances For
        @[simp]
        theorem Subring.coe_map {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (s : Subring R) :
        (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 map f s xs, f x = y
        @[simp]
        theorem Subring.map_id {R : Type u} [Ring R] (s : Subring R) :
        map (RingHom.id R) s = s
        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) :
        map g (map f s) = map (g.comp f) 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} :
        map f s t s comap f t
        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) :
        s ≃+* (map f s)

        A subring is isomorphic to its image under an injective function

        Equations
        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 : s) :
          ((s.equivMapOfInjective 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) :
            f.range = Set.range f
            @[simp]
            theorem RingHom.mem_range {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} {y : S} :
            y f.range ∃ (x : R), 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) :
            f x f.range
            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.instBot {R : Type u} [Ring R] :
            Equations
            instance Subring.instInhabited {R : Type u} [Ring R] :
            Equations
            theorem Subring.mem_bot {R : Type u} [Ring R] {x : R} :
            x ∃ (n : ), n = x

            inf #

            instance Subring.instMin {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 p' : Subring R) :
            ↑(p p') = p p'
            @[simp]
            theorem Subring.mem_inf {R : Type u} [Ring R] {p p' : Subring R} {x : R} :
            x p p' x p x p'
            instance Subring.instInfSet {R : Type u} [Ring R] :
            Equations
            @[simp]
            theorem Subring.coe_sInf {R : Type u} [Ring R] (S : Set (Subring R)) :
            (sInf S) = sS, s
            theorem Subring.mem_sInf {R : Type u} [Ring R] {S : Set (Subring R)} {x : R} :
            x sInf S pS, x 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 = ts, t.toSubmonoid
            @[simp]
            theorem Subring.sInf_toAddSubgroup {R : Type u} [Ring R] (s : Set (Subring R)) :
            (sInf s).toAddSubgroup = ts, t.toAddSubgroup

            Subrings of a ring form a complete lattice.

            Equations
            theorem Subring.eq_top_iff' {R : Type u} [Ring R] (A : Subring R) :
            A = ∀ (x : R), x A

            Center of a ring #

            def Subring.center (R : Type u) [Ring R] :

            The center of a ring R is the set of elements that commute with everything in R

            Equations
            • Subring.center R = { carrier := Set.center R, mul_mem' := , one_mem' := , add_mem' := , zero_mem' := , neg_mem' := }
            Instances For
              theorem Subring.coe_center (R : Type u) [Ring R] :
              theorem Subring.mem_center_iff {R : Type u} [Ring R] {z : R} :
              z center R ∀ (g : R), g * z = z * g
              instance Subring.decidableMemCenter {R : Type u} [Ring R] [DecidableEq R] [Fintype R] :
              DecidablePred fun (x : R) => x center R
              Equations
              @[simp]
              theorem Subring.center_eq_top (R : Type u_1) [CommRing R] :
              def Subring.centerCongr {R : Type u} {S : Type v} [Ring R] [Ring S] (e : R ≃+* S) :
              (center R) ≃+* (center S)

              The center of isomorphic (not necessarily associative) rings are isomorphic.

              Equations
              Instances For
                @[simp]
                theorem Subring.centerCongr_symm_apply_coe {R : Type u} {S : Type v} [Ring R] [Ring S] (e : R ≃+* S) (s : (Subsemigroup.center S)) :
                ((centerCongr e).symm s) = (↑e).symm s
                @[simp]
                theorem Subring.centerCongr_apply_coe {R : Type u} {S : Type v} [Ring R] [Ring S] (e : R ≃+* S) (r : (Subsemigroup.center R)) :
                ((centerCongr e) r) = e r

                The center of a (not necessarily associative) ring is isomorphic to the center of its opposite.

                Equations
                Instances For
                  instance Subring.instField {K : Type u} [DivisionRing K] :
                  Field (center K)
                  Equations
                  @[simp]
                  theorem Subring.center.coe_inv {K : Type u} [DivisionRing K] (a : (center K)) :
                  a⁻¹ = (↑a)⁻¹
                  @[simp]
                  theorem Subring.center.coe_div {K : Type u} [DivisionRing K] (a b : (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
                  Instances For
                    @[simp]
                    theorem Subring.coe_centralizer {R : Type u} [Ring R] (s : Set R) :
                    theorem Subring.mem_centralizer_iff {R : Type u} [Ring R] {s : Set R} {z : R} :
                    z centralizer s gs, g * z = z * g
                    theorem Subring.centralizer_le {R : Type u} [Ring R] (s t : Set R) (h : s t) :
                    @[simp]

                    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 closure s ∀ (S : Subring R), s Sx S
                      @[simp]
                      theorem Subring.subset_closure {R : Type u} [Ring R] {s : Set R} :
                      s (closure s)

                      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 : Pclosure s) :
                      Ps
                      @[simp]
                      theorem Subring.closure_le {R : Type u} [Ring R] {s : Set R} {t : Subring R} :
                      closure s t s t

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

                      theorem Subring.closure_mono {R : Type u} [Ring R] ⦃s 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 closure s) :
                      theorem Subring.closure_induction {R : Type u} [Ring R] {s : Set R} {p : (x : R) → x closure sProp} (mem : ∀ (x : R) (hx : x s), p x ) (zero : p 0 ) (one : p 1 ) (add : ∀ (x y : R) (hx : x closure s) (hy : y closure s), p x hxp y hyp (x + y) ) (neg : ∀ (x : R) (hx : x closure s), p x hxp (-x) ) (mul : ∀ (x y : R) (hx : x closure s) (hy : y closure s), p x hxp y hyp (x * y) ) {x : R} (hx : x closure s) :
                      p x hx

                      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.

                      @[deprecated Subring.closure_induction (since := "2024-10-10")]
                      theorem Subring.closure_induction' {R : Type u} [Ring R] {s : Set R} {p : (x : R) → x closure sProp} (mem : ∀ (x : R) (hx : x s), p x ) (zero : p 0 ) (one : p 1 ) (add : ∀ (x y : R) (hx : x closure s) (hy : y closure s), p x hxp y hyp (x + y) ) (neg : ∀ (x : R) (hx : x closure s), p x hxp (-x) ) (mul : ∀ (x y : R) (hx : x closure s) (hy : y closure s), p x hxp y hyp (x * y) ) {x : R} (hx : x closure s) :
                      p x hx

                      Alias of Subring.closure_induction.


                      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 y : R) → x closure sy closure sProp} (mem_mem : ∀ (x y : R) (hx : x s) (hy : y s), p x y ) (zero_left : ∀ (x : R) (hx : x closure s), p 0 x hx) (zero_right : ∀ (x : R) (hx : x closure s), p x 0 hx ) (one_left : ∀ (x : R) (hx : x closure s), p 1 x hx) (one_right : ∀ (x : R) (hx : x closure s), p x 1 hx ) (neg_left : ∀ (x y : R) (hx : x closure s) (hy : y closure s), p x y hx hyp (-x) y hy) (neg_right : ∀ (x y : R) (hx : x closure s) (hy : y closure s), p x y hx hyp x (-y) hx ) (add_left : ∀ (x y z : R) (hx : x closure s) (hy : y closure s) (hz : z closure s), p x z hx hzp y z hy hzp (x + y) z hz) (add_right : ∀ (x y z : R) (hx : x closure s) (hy : y closure s) (hz : z closure s), p x y hx hyp x z hx hzp x (y + z) hx ) (mul_left : ∀ (x y z : R) (hx : x closure s) (hy : y closure s) (hz : z closure s), p x z hx hzp y z hy hzp (x * y) z hz) (mul_right : ∀ (x y z : R) (hx : x closure s) (hy : y closure s) (hz : z closure s), p x y hx hyp x z hx hzp x (y * z) hx ) {x y : R} (hx : x closure s) (hy : y closure s) :
                      p x y hx hy

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

                      @[reducible, inline]
                      abbrev Subring.closureCommRingOfComm {R : Type u} [Ring R] {s : Set R} (hcomm : xs, ys, x * y = y * x) :

                      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} (hx : x closure s) :
                        ∃ (L : List (List R)), (∀ tL, yt, y s y = -1) (List.map List.prod L).sum = x

                        closure forms a Galois insertion with the coercion to set.

                        Equations
                        Instances For
                          @[simp]
                          theorem Subring.closure_eq {R : Type u} [Ring R] (s : Subring R) :
                          closure s = s

                          Closure of a subring S equals S.

                          @[simp]
                          theorem Subring.closure_union {R : Type u} [Ring R] (s t : Set R) :
                          theorem Subring.closure_iUnion {R : Type u} [Ring R] {ι : Sort u_1} (s : ιSet R) :
                          closure (⋃ (i : ι), s i) = ⨆ (i : ι), closure (s i)
                          theorem Subring.closure_sUnion {R : Type u} [Ring R] (s : Set (Set R)) :
                          closure (⋃₀ s) = ts, closure t
                          theorem Subring.map_sup {R : Type u} {S : Type v} [Ring R] [Ring S] (s t : Subring R) (f : R →+* S) :
                          map f (s t) = map f s map f t
                          theorem Subring.map_iSup {R : Type u} {S : Type v} [Ring R] [Ring S] {ι : Sort u_1} (f : R →+* S) (s : ιSubring R) :
                          map f (iSup s) = ⨆ (i : ι), map f (s i)
                          theorem Subring.map_inf {R : Type u} {S : Type v} [Ring R] [Ring S] (s t : Subring R) (f : R →+* S) (hf : Function.Injective f) :
                          map f (s t) = map f s map f t
                          theorem Subring.map_iInf {R : Type u} {S : Type v} [Ring R] [Ring S] {ι : Sort u_1} [Nonempty ι] (f : R →+* S) (hf : Function.Injective f) (s : ιSubring R) :
                          map f (iInf s) = ⨅ (i : ι), map f (s i)
                          theorem Subring.comap_inf {R : Type u} {S : Type v} [Ring R] [Ring S] (s t : Subring S) (f : R →+* S) :
                          comap f (s t) = comap f s comap f t
                          theorem Subring.comap_iInf {R : Type u} {S : Type v} [Ring R] [Ring S] {ι : Sort u_1} (f : R →+* S) (s : ιSubring S) :
                          comap f (iInf s) = ⨅ (i : ι), 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
                          • s.prod t = { carrier := s ×ˢ t, mul_mem' := , one_mem' := , add_mem' := , zero_mem' := , neg_mem' := }
                          Instances For
                            theorem Subring.coe_prod {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) (t : Subring S) :
                            (s.prod 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 s.prod t p.1 s p.2 t
                            theorem Subring.prod_mono {R : Type u} {S : Type v} [Ring R] [Ring S] ⦃s₁ s₂ : Subring R (hs : s₁ s₂) ⦃t₁ t₂ : Subring S (ht : t₁ t₂) :
                            s₁.prod t₁ s₂.prod t₂
                            theorem Subring.prod_mono_right {R : Type u} {S : Type v} [Ring R] [Ring S] (s : Subring R) :
                            Monotone fun (t : Subring S) => s.prod t
                            theorem Subring.prod_mono_left {R : Type u} {S : Type v} [Ring R] [Ring S] (t : Subring S) :
                            Monotone fun (s : Subring R) => s.prod 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) :
                            (s.prod t) ≃+* s × t

                            Product of subrings is isomorphic to their product as rings.

                            Equations
                            Instances For
                              theorem Subring.mem_iSup_of_directed {R : Type u} [Ring R] {ι : Sort u_1} [hι : Nonempty ι] {S : ιSubring R} (hS : Directed (fun (x1 x2 : Subring R) => x1 x2) 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 (x1 x2 : Subring R) => x1 x2) S) :
                              (⨆ (i : ι), S i) = ⋃ (i : ι), (S i)
                              theorem Subring.mem_sSup_of_directedOn {R : Type u} [Ring R] {S : Set (Subring R)} (Sne : S.Nonempty) (hS : DirectedOn (fun (x1 x2 : Subring R) => x1 x2) S) {x : R} :
                              x sSup S sS, x s
                              theorem Subring.coe_sSup_of_directedOn {R : Type u} [Ring R] {S : Set (Subring R)} (Sne : S.Nonempty) (hS : DirectedOn (fun (x1 x2 : Subring R) => x1 x2) S) :
                              (sSup S) = sS, 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 map (↑f) K f.symm 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) :
                              map (↑f) K = comap (↑f.symm) K
                              theorem Subring.comap_equiv_eq_map_symm {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R ≃+* S) (K : Subring S) :
                              comap (↑f) K = map (↑f.symm) K
                              def RingHom.rangeRestrict {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) :
                              R →+* f.range

                              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) :
                                (f.rangeRestrict x) = f x
                                theorem RingHom.range_eq_top {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} :
                                @[deprecated RingHom.range_eq_top (since := "2024-11-11")]
                                theorem RingHom.range_top_iff_surjective {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} :

                                Alias of RingHom.range_eq_top.

                                @[simp]
                                theorem RingHom.range_eq_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.

                                @[deprecated RingHom.range_eq_top_of_surjective (since := "2024-11-11")]
                                theorem RingHom.range_top_of_surjective {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (hf : Function.Surjective f) :

                                Alias of RingHom.range_eq_top_of_surjective.


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

                                def RingHom.eqLocus {R : Type u} [Ring R] {S : Type v} [Semiring S] (f 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
                                • f.eqLocus g = { carrier := {x : R | f x = g x}, mul_mem' := , one_mem' := , add_mem' := , zero_mem' := , neg_mem' := }
                                Instances For
                                  @[simp]
                                  theorem RingHom.eqLocus_same {R : Type u} [Ring R] {S : Type v} [Semiring S] (f : R →+* S) :
                                  theorem RingHom.eqOn_set_closure {R : Type u} [Ring R] {S : Type v} [Semiring S] {f 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} [Ring R] {S : Type v} [Semiring S] {f g : R →+* S} (h : Set.EqOn f g ) :
                                  f = g
                                  theorem RingHom.eq_of_eqOn_set_dense {R : Type u} [Ring R] {S : Type v} [Semiring S] {s : Set R} (hs : Subring.closure s = ) {f 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 T : Subring R} (h : S T) :
                                  S →+* T

                                  The ring homomorphism associated to an inclusion of subrings.

                                  Equations
                                  Instances For
                                    @[simp]
                                    theorem Subring.range_subtype {R : Type u} [Ring R] (s : Subring R) :
                                    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 t : Subring R} (h : s = t) :
                                    s ≃+* t

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

                                    Equations
                                    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 ≃+* f.range

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

                                      Equations
                                      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) :
                                        ((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 : f.range) :
                                        (ofLeftInverse h).symm x = g x
                                        def RingEquiv.subringMap {R : Type u} {S : Type v} [Ring R] [Ring S] {s : Subring R} (e : R ≃+* 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 closure s) (h1 : C 1) (hneg1 : C (-1)) (hs : zs, ∀ (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) :
                                          closure (f ⁻¹' s) comap f (closure s)

                                          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.instSMulSubtypeMem {R : Type u} [Ring R] {α : Type u_1} [SMul R α] (S : Subring R) :
                                          SMul (↥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 : 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 (↥S) α β
                                          instance Subring.smulCommClass_right {R : Type u} [Ring R] {α : Type u_1} {β : Type u_2} [SMul α β] [SMul R β] [SMulCommClass α R β] (S : Subring R) :
                                          SMulCommClass α (↥S) β
                                          instance Subring.instIsScalarTowerSubtypeMem {R : Type u} [Ring R] {α : Type u_1} {β : Type u_2} [SMul α β] [SMul R α] [SMul R β] [IsScalarTower R α β] (S : Subring R) :
                                          IsScalarTower (↥S) α β

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

                                          instance Subring.instFaithfulSMulSubtypeMem {R : Type u} [Ring R] {α : Type u_1} [SMul R α] [FaithfulSMul R α] (S : Subring R) :
                                          FaithfulSMul (↥S) α
                                          instance Subring.instMulActionSubtypeMem {R : Type u} [Ring R] {α : Type u_1} [MulAction R α] (S : Subring R) :
                                          MulAction (↥S) α

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

                                          Equations
                                          instance Subring.instDistribMulActionSubtypeMem {R : Type u} [Ring R] {α : Type u_1} [AddMonoid α] [DistribMulAction R α] (S : Subring R) :

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

                                          Equations
                                          instance Subring.instMulDistribMulActionSubtypeMem {R : Type u} [Ring R] {α : Type u_1} [Monoid α] [MulDistribMulAction R α] (S : Subring R) :

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

                                          Equations
                                          instance Subring.instSMulWithZeroSubtypeMem {R : Type u} [Ring R] {α : Type u_1} [Zero α] [SMulWithZero R α] (S : Subring R) :
                                          SMulWithZero (↥S) α

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

                                          Equations
                                          instance Subring.instMulActionWithZeroSubtypeMem {R : Type u} [Ring R] {α : Type u_1} [Zero α] [MulActionWithZero R α] (S : Subring R) :

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

                                          Equations
                                          instance Subring.instModuleSubtypeMem {R : Type u} [Ring R] {α : Type u_1} [AddCommMonoid α] [Module R α] (S : Subring R) :
                                          Module (↥S) α

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

                                          Equations
                                          instance Subring.instMulSemiringActionSubtypeMem {R : Type u} [Ring R] {α : Type u_1} [Semiring α] [MulSemiringAction R α] (S : Subring R) :

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

                                          Equations

                                          The center of a semiring acts commutatively on that semiring.

                                          The center of a semiring acts commutatively on that semiring.

                                          theorem Subring.map_comap_eq {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (t : Subring S) :
                                          map f (comap f t) = t f.range
                                          theorem Subring.map_comap_eq_self {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} {t : Subring S} (h : t f.range) :
                                          map f (comap f t) = t
                                          theorem Subring.map_comap_eq_self_of_surjective {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} (hf : Function.Surjective f) (t : Subring S) :
                                          map f (comap f t) = t
                                          theorem Subring.comap_map_eq {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) (s : Subring R) :
                                          comap f (map f s) = s closure (f ⁻¹' {0})
                                          theorem Subring.comap_map_eq_self {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} {s : Subring R} (h : f ⁻¹' {0} s) :
                                          comap f (map f s) = s
                                          theorem Subring.comap_map_eq_self_of_injective {R : Type u} {S : Type v} [Ring R] [Ring S] {f : R →+* S} (hf : Function.Injective f) (s : Subring R) :
                                          comap f (map f s) = s
                                          theorem AddSubgroup.int_mul_mem {R : Type u} [Ring R] {G : AddSubgroup R} (k : ) {g : R} (h : g G) :
                                          k * g G