Documentation

Mathlib.RingTheory.NonUnitalSubring.Basic

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)

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

theorem NonUnitalSubring.list_sum_mem {R : Type u} [NonUnitalNonAssocRing R] (s : NonUnitalSubring R) {l : List R} :
(∀ xl, x s)l.sum s

Sum of a list of elements in a non-unital subring is in the non-unital subring.

theorem NonUnitalSubring.multiset_sum_mem {R : Type u_1} [NonUnitalNonAssocRing R] (s : NonUnitalSubring R) (m : Multiset R) :
(∀ am, a s)m.sum s

Sum of a multiset of elements in a NonUnitalSubring of a NonUnitalRing is in the NonUnitalSubring.

theorem NonUnitalSubring.sum_mem {R : Type u_1} [NonUnitalNonAssocRing R] (s : NonUnitalSubring R) {ι : Type u_2} {t : Finset ι} {f : ιR} (h : ct, f c s) :
it, f i s

Sum of elements in a NonUnitalSubring of a NonUnitalRing indexed by a Finset is in the NonUnitalSubring.

top #

The non-unital subring R of the ring R.

Equations
  • NonUnitalSubring.instTop = { top := let __src := ; let __src_1 := ; { carrier := __src.carrier, add_mem' := , zero_mem' := , mul_mem' := , neg_mem' := } }
@[simp]
@[simp]
theorem NonUnitalSubring.coe_top {R : Type u} [NonUnitalNonAssocRing R] :
= Set.univ

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

Equations
  • NonUnitalSubring.topEquiv = NonUnitalSubsemiring.topEquiv
Instances For
    @[simp]
    theorem NonUnitalSubring.topEquiv_apply {R : Type u} [NonUnitalNonAssocRing R] (x : ) :
    NonUnitalSubring.topEquiv x = x
    @[simp]
    theorem NonUnitalSubring.topEquiv_symm_apply_coe {R : Type u} [NonUnitalNonAssocRing R] (x : R) :
    (NonUnitalSubring.topEquiv.symm x) = x

    comap #

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

    Equations
    Instances For
      @[simp]
      theorem NonUnitalSubring.coe_comap {F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] (s : NonUnitalSubring S) (f : F) :
      (NonUnitalSubring.comap f s) = f ⁻¹' s
      @[simp]
      theorem NonUnitalSubring.mem_comap {F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] {s : NonUnitalSubring S} {f : F} {x : R} :

      map #

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

      Equations
      • NonUnitalSubring.map f s = { carrier := f '' s.carrier, add_mem' := , zero_mem' := , mul_mem' := , neg_mem' := }
      Instances For
        @[simp]
        theorem NonUnitalSubring.coe_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.map f s) = f '' s
        @[simp]
        theorem NonUnitalSubring.mem_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} {y : S} :
        y NonUnitalSubring.map f s xs, f x = y
        noncomputable def NonUnitalSubring.equivMapOfInjective {F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] (s : NonUnitalSubring R) (f : F) (hf : Function.Injective f) :

        A NonUnitalSubring is isomorphic to its image under an injective function

        Equations
        • s.equivMapOfInjective f hf = { toEquiv := Equiv.Set.image (⇑f) (↑s) hf, map_mul' := , map_add' := }
        Instances For
          @[simp]
          theorem NonUnitalSubring.coe_equivMapOfInjective_apply {F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] (s : NonUnitalSubring R) (f : F) (hf : Function.Injective f) (x : s) :
          ((s.equivMapOfInjective f hf) x) = f x

          range #

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

          Equations
          Instances For
            @[simp]
            theorem NonUnitalRingHom.coe_range {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] (f : R →ₙ+* S) :
            f.range = Set.range f
            @[simp]
            theorem NonUnitalRingHom.mem_range {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] {f : R →ₙ+* S} {y : S} :
            y f.range ∃ (x : R), f x = y
            theorem NonUnitalRingHom.mem_range_self {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] (f : R →ₙ+* S) (x : R) :
            f x f.range
            theorem NonUnitalRingHom.map_range {R : Type u} {S : Type v} {T : Type u_1} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [NonUnitalNonAssocRing T] (g : S →ₙ+* T) (f : R →ₙ+* S) :
            NonUnitalSubring.map g f.range = (g.comp f).range

            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 #

            Equations
            Equations
            • NonUnitalSubring.instInhabited = { default := }

            inf #

            The inf of two NonUnitalSubrings is their intersection.

            Equations
            • One or more equations did not get rendered due to their size.
            @[simp]
            theorem NonUnitalSubring.coe_inf {R : Type u} [NonUnitalNonAssocRing R] (p : NonUnitalSubring R) (p' : NonUnitalSubring R) :
            (p p') = p p'
            @[simp]
            theorem NonUnitalSubring.mem_inf {R : Type u} [NonUnitalNonAssocRing R] {p : NonUnitalSubring R} {p' : NonUnitalSubring R} {x : R} :
            x p p' x p x p'
            Equations
            @[simp]
            theorem NonUnitalSubring.coe_sInf {R : Type u} [NonUnitalNonAssocRing R] (S : Set (NonUnitalSubring R)) :
            (sInf S) = sS, s
            theorem NonUnitalSubring.mem_sInf {R : Type u} [NonUnitalNonAssocRing R] {S : Set (NonUnitalSubring R)} {x : R} :
            x sInf S pS, x p
            @[simp]
            theorem NonUnitalSubring.coe_iInf {R : Type u} [NonUnitalNonAssocRing R] {ι : Sort u_1} {S : ιNonUnitalSubring R} :
            (⨅ (i : ι), S i) = ⋂ (i : ι), (S i)
            theorem NonUnitalSubring.mem_iInf {R : Type u} [NonUnitalNonAssocRing R] {ι : Sort u_1} {S : ιNonUnitalSubring R} {x : R} :
            x ⨅ (i : ι), S i ∀ (i : ι), x S i
            @[simp]
            theorem NonUnitalSubring.sInf_toSubsemigroup {R : Type u} [NonUnitalNonAssocRing R] (s : Set (NonUnitalSubring R)) :
            (sInf s).toSubsemigroup = ts, t.toSubsemigroup
            @[simp]
            theorem NonUnitalSubring.sInf_toAddSubgroup {R : Type u} [NonUnitalNonAssocRing R] (s : Set (NonUnitalSubring R)) :
            (sInf s).toAddSubgroup = ts, t.toAddSubgroup

            NonUnitalSubrings of a ring form a complete lattice.

            Equations

            Center of a ring #

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

            Equations
            Instances For
              theorem NonUnitalSubring.mem_center_iff {R : Type u} [NonUnitalRing R] {z : R} :
              z NonUnitalSubring.center R ∀ (g : R), g * z = z * g

              NonUnitalSubring closure of a subset #

              The NonUnitalSubring generated by a set.

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

                The NonUnitalSubring generated by a set includes the set.

                @[simp]

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

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

                theorem NonUnitalSubring.closure_induction {R : Type u} [NonUnitalNonAssocRing R] {s : Set R} {p : (x : R) → x NonUnitalSubring.closure sProp} (mem : ∀ (x : R) (hx : x s), p x ) (zero : p 0 ) (add : ∀ (x y : R) (hx : x NonUnitalSubring.closure s) (hy : y NonUnitalSubring.closure s), p x hxp y hyp (x + y) ) (neg : ∀ (x : R) (hx : x NonUnitalSubring.closure s), p x hxp (-x) ) (mul : ∀ (x y : R) (hx : x NonUnitalSubring.closure s) (hy : y NonUnitalSubring.closure s), p x hxp y hyp (x * y) ) {x : R} (hx : x NonUnitalSubring.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 NonUnitalSubring.closure_induction]
                theorem NonUnitalSubring.closure_induction' {R : Type u} [NonUnitalNonAssocRing R] {s : Set R} {p : (NonUnitalSubring.closure s)Prop} (a : (NonUnitalSubring.closure s)) (mem : ∀ (x : R) (hx : x s), p x, ) (zero : p 0) (add : ∀ (x y : (NonUnitalSubring.closure s)), p xp yp (x + y)) (neg : ∀ (x : (NonUnitalSubring.closure s)), p xp (-x)) (mul : ∀ (x y : (NonUnitalSubring.closure s)), p xp yp (x * y)) :
                p a

                The difference with NonUnitalSubring.closure_induction is that this acts on the subtype.

                theorem NonUnitalSubring.closure_induction₂ {R : Type u} [NonUnitalNonAssocRing R] {s : Set R} {p : (x y : R) → x NonUnitalSubring.closure sy NonUnitalSubring.closure sProp} (mem_mem : ∀ (x y : R) (hx : x s) (hy : y s), p x y ) (zero_left : ∀ (x : R) (hx : x NonUnitalSubring.closure s), p 0 x hx) (zero_right : ∀ (x : R) (hx : x NonUnitalSubring.closure s), p x 0 hx ) (neg_left : ∀ (x y : R) (hx : x NonUnitalSubring.closure s) (hy : y NonUnitalSubring.closure s), p x y hx hyp (-x) y hy) (neg_right : ∀ (x y : R) (hx : x NonUnitalSubring.closure s) (hy : y NonUnitalSubring.closure s), p x y hx hyp x (-y) hx ) (add_left : ∀ (x y z : R) (hx : x NonUnitalSubring.closure s) (hy : y NonUnitalSubring.closure s) (hz : z NonUnitalSubring.closure s), p x z hx hzp y z hy hzp (x + y) z hz) (add_right : ∀ (x y z : R) (hx : x NonUnitalSubring.closure s) (hy : y NonUnitalSubring.closure s) (hz : z NonUnitalSubring.closure s), p x y hx hyp x z hx hzp x (y + z) hx ) (mul_left : ∀ (x y z : R) (hx : x NonUnitalSubring.closure s) (hy : y NonUnitalSubring.closure s) (hz : z NonUnitalSubring.closure s), p x z hx hzp y z hy hzp (x * y) z hz) (mul_right : ∀ (x y z : R) (hx : x NonUnitalSubring.closure s) (hy : y NonUnitalSubring.closure s) (hz : z NonUnitalSubring.closure s), p x y hx hyp x z hx hzp x (y * z) hx ) {x : R} {y : R} (hx : x NonUnitalSubring.closure s) (hy : y NonUnitalSubring.closure s) :
                p x y hx hy

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

                def NonUnitalSubring.closureNonUnitalCommRingOfComm {R : Type u} [NonUnitalRing R] {s : Set R} (hcomm : as, bs, a * b = b * a) :

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

                Equations
                Instances For
                  def NonUnitalSubring.gi (R : Type u) [NonUnitalNonAssocRing R] :
                  GaloisInsertion NonUnitalSubring.closure SetLike.coe

                  closure forms a Galois insertion with the coercion to set.

                  Equations
                  Instances For
                    theorem NonUnitalSubring.closure_iUnion {R : Type u} [NonUnitalNonAssocRing R] {ι : Sort u_1} (s : ιSet R) :
                    NonUnitalSubring.closure (⋃ (i : ι), s i) = ⨆ (i : ι), NonUnitalSubring.closure (s i)
                    theorem NonUnitalSubring.map_iSup {F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] {ι : Sort u_1} (f : F) (s : ιNonUnitalSubring R) :
                    NonUnitalSubring.map f (iSup s) = ⨆ (i : ι), NonUnitalSubring.map f (s i)
                    theorem NonUnitalSubring.map_iInf {F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] {ι : Sort u_1} [Nonempty ι] (f : F) (hf : Function.Injective f) (s : ιNonUnitalSubring R) :
                    NonUnitalSubring.map f (iInf s) = ⨅ (i : ι), NonUnitalSubring.map f (s i)
                    theorem NonUnitalSubring.comap_iInf {F : Type w} {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] [FunLike F R S] [NonUnitalRingHomClass F R S] {ι : Sort u_1} (f : F) (s : ιNonUnitalSubring S) :

                    Given NonUnitalSubrings s, t of rings R, S respectively, s.prod t is s ×ˢ t as a NonUnitalSubring of R × S.

                    Equations
                    • s.prod t = { carrier := s ×ˢ t, add_mem' := , zero_mem' := , mul_mem' := , neg_mem' := }
                    Instances For
                      theorem NonUnitalSubring.coe_prod {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] (s : NonUnitalSubring R) (t : NonUnitalSubring S) :
                      (s.prod t) = s ×ˢ t
                      theorem NonUnitalSubring.mem_prod {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] {s : NonUnitalSubring R} {t : NonUnitalSubring S} {p : R × S} :
                      p s.prod t p.1 s p.2 t
                      theorem NonUnitalSubring.prod_mono {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] ⦃s₁ : NonUnitalSubring R ⦃s₂ : NonUnitalSubring R (hs : s₁ s₂) ⦃t₁ : NonUnitalSubring S ⦃t₂ : NonUnitalSubring S (ht : t₁ t₂) :
                      s₁.prod t₁ s₂.prod t₂
                      def NonUnitalSubring.prodEquiv {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] (s : NonUnitalSubring R) (t : NonUnitalSubring S) :
                      (s.prod t) ≃+* s × t

                      Product of NonUnitalSubrings is isomorphic to their product as rings.

                      Equations
                      • s.prodEquiv t = { toEquiv := Equiv.Set.prod s t, map_mul' := , map_add' := }
                      Instances For
                        theorem NonUnitalSubring.mem_iSup_of_directed {R : Type u} [NonUnitalNonAssocRing R] {ι : Sort u_1} [hι : Nonempty ι] {S : ιNonUnitalSubring R} (hS : Directed (fun (x1 x2 : NonUnitalSubring R) => x1 x2) S) {x : R} :
                        x ⨆ (i : ι), S i ∃ (i : ι), x S i

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

                        theorem NonUnitalSubring.coe_iSup_of_directed {R : Type u} [NonUnitalNonAssocRing R] {ι : Sort u_1} [Nonempty ι] {S : ιNonUnitalSubring R} (hS : Directed (fun (x1 x2 : NonUnitalSubring R) => x1 x2) S) :
                        (⨆ (i : ι), S i) = ⋃ (i : ι), (S i)
                        theorem NonUnitalSubring.mem_sSup_of_directedOn {R : Type u} [NonUnitalNonAssocRing R] {S : Set (NonUnitalSubring R)} (Sne : S.Nonempty) (hS : DirectedOn (fun (x1 x2 : NonUnitalSubring R) => x1 x2) S) {x : R} :
                        x sSup S sS, x s
                        theorem NonUnitalSubring.coe_sSup_of_directedOn {R : Type u} [NonUnitalNonAssocRing R] {S : Set (NonUnitalSubring R)} (Sne : S.Nonempty) (hS : DirectedOn (fun (x1 x2 : NonUnitalSubring R) => x1 x2) S) :
                        (sSup S) = sS, s
                        theorem NonUnitalSubring.mem_map_equiv {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] {f : R ≃+* S} {K : NonUnitalSubring R} {x : S} :
                        x NonUnitalSubring.map (↑f) K f.symm x K

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

                        This is the bundled version of Set.rangeFactorization.

                        Equations
                        Instances For
                          @[simp]
                          theorem NonUnitalRingHom.coe_rangeRestrict {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] (f : R →ₙ+* S) (x : R) :
                          (f.rangeRestrict x) = f x
                          @[simp]

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

                          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

                          Equations
                          • f.eqLocus g = { carrier := {x : R | f x = g x}, add_mem' := , zero_mem' := , mul_mem' := , neg_mem' := }
                          Instances For
                            @[simp]
                            theorem NonUnitalRingHom.eqLocus_same {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] (f : R →ₙ+* S) :
                            f.eqLocus f =
                            theorem NonUnitalRingHom.eqOn_set_closure {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] {f : R →ₙ+* S} {g : R →ₙ+* S} {s : Set R} (h : Set.EqOn (⇑f) (⇑g) s) :

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

                            theorem NonUnitalRingHom.eq_of_eqOn_set_top {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] {f : R →ₙ+* S} {g : R →ₙ+* S} (h : Set.EqOn f g ) :
                            f = g
                            theorem NonUnitalRingHom.eq_of_eqOn_set_dense {R : Type u} {S : Type v} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] {s : Set R} (hs : NonUnitalSubring.closure s = ) {f : R →ₙ+* S} {g : R →ₙ+* S} (h : Set.EqOn (⇑f) (⇑g) s) :
                            f = g

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

                            def RingEquiv.nonUnitalSubringCongr {R : Type u} [NonUnitalRing R] {s : NonUnitalSubring R} {t : NonUnitalSubring R} (h : s = t) :
                            s ≃+* t

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

                            Equations
                            Instances For
                              def RingEquiv.ofLeftInverse' {R : Type u} {S : Type v} [NonUnitalRing R] [NonUnitalRing 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
                              • 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} [NonUnitalRing R] [NonUnitalRing S] {g : SR} {f : R →ₙ+* S} (h : Function.LeftInverse g f) (x : R) :
                                @[simp]
                                theorem RingEquiv.ofLeftInverse'_symm_apply {R : Type u} {S : Type v} [NonUnitalRing R] [NonUnitalRing S] {g : SR} {f : R →ₙ+* S} (h : Function.LeftInverse g f) (x : f.range) :
                                (RingEquiv.ofLeftInverse' h).symm x = g x