Documentation

Mathlib.GroupTheory.Submonoid.Basic

Submonoids: definition and CompleteLattice structure #

This file defines bundled multiplicative and additive submonoids. We also define a CompleteLattice structure on Submonoids, define the closure of a set as the minimal submonoid that includes this set, and prove a few results about extending properties from a dense set (i.e. a set with closure s = ⊤) to the whole monoid, see Submonoid.dense_induction and MonoidHom.ofClosureEqTopLeft/MonoidHom.ofClosureEqTopRight.

Main definitions #

For each of the following definitions in the Submonoid namespace, there is a corresponding definition in the AddSubmonoid namespace.

Implementation notes #

Submonoid inclusion is denoted rather than , although is defined as membership of a submonoid's underlying set.

Note that Submonoid M does not actually require Monoid M, instead requiring only the weaker MulOneClass M.

This file is designed to have very few dependencies. In particular, it should not use natural numbers. Submonoid is implemented by extending Subsemigroup requiring one_mem'.

Tags #

submonoid, submonoids

class OneMemClass (S : Type u_4) (M : Type u_5) [One M] [SetLike S M] :
  • one_mem : ∀ (s : S), 1 s

    By definition, if we have OneMemClass S M, we have 1 ∈ s for all s : S.

OneMemClass S M says S is a type of subsets s ≤ M, such that 1 ∈ s for all s.

Instances
    class ZeroMemClass (S : Type u_4) (M : Type u_5) [Zero M] [SetLike S M] :
    • zero_mem : ∀ (s : S), 0 s

      By definition, if we have ZeroMemClass S M, we have 0 ∈ s for all s : S.

    ZeroMemClass S M says S is a type of subsets s ≤ M, such that 0 ∈ s for all s.

    Instances
      structure Submonoid (M : Type u_4) [MulOneClass M] extends Subsemigroup :
      Type u_4
      • carrier : Set M
      • mul_mem' : ∀ {a b : M}, a s.carrierb s.carriera * b s.carrier
      • one_mem' : 1 s.carrier

        A submonoid contains 1.

      A submonoid of a monoid M is a subset containing 1 and closed under multiplication.

      Instances For
        class SubmonoidClass (S : Type u_4) (M : Type u_5) [MulOneClass M] [SetLike S M] extends MulMemClass , OneMemClass :

          SubmonoidClass S M says S is a type of subsets s ≤ M that contain 1 and are closed under (*)

          Instances
            structure AddSubmonoid (M : Type u_4) [AddZeroClass M] extends AddSubsemigroup :
            Type u_4
            • carrier : Set M
            • add_mem' : ∀ {a b : M}, a s.carrierb s.carriera + b s.carrier
            • zero_mem' : 0 s.carrier

              An additive submonoid contains 0.

            An additive submonoid of an additive monoid M is a subset containing 0 and closed under addition.

            Instances For
              class AddSubmonoidClass (S : Type u_4) (M : Type u_5) [AddZeroClass M] [SetLike S M] extends AddMemClass , ZeroMemClass :

                AddSubmonoidClass S M says S is a type of subsets s ≤ M that contain 0 and are closed under (+)

                Instances
                  theorem nsmul_mem {M : Type u_4} {A : Type u_5} [AddMonoid M] [SetLike A M] [AddSubmonoidClass A M] {S : A} {x : M} (hx : x S) (n : ) :
                  n x S
                  abbrev nsmul_mem.match_1 (motive : Prop) :
                  (x : ) → (Unitmotive 0) → ((n : ) → motive (Nat.succ n)) → motive x
                  Equations
                  Instances For
                    theorem pow_mem {M : Type u_4} {A : Type u_5} [Monoid M] [SetLike A M] [SubmonoidClass A M] {S : A} {x : M} (hx : x S) (n : ) :
                    x ^ n S
                    theorem AddSubmonoid.instSetLikeAddSubmonoid.proof_1 {M : Type u_1} [AddZeroClass M] (p : AddSubmonoid M) (q : AddSubmonoid M) (h : (fun s => s.carrier) p = (fun s => s.carrier) q) :
                    p = q
                    Equations
                    • AddSubmonoid.instSetLikeAddSubmonoid = { coe := fun s => s.carrier, coe_injective' := (_ : ∀ (p q : AddSubmonoid M), (fun s => s.carrier) p = (fun s => s.carrier) qp = q) }
                    Equations
                    • Submonoid.instSetLikeSubmonoid = { coe := fun s => s.carrier, coe_injective' := (_ : ∀ (p q : Submonoid M), (fun s => s.carrier) p = (fun s => s.carrier) qp = q) }
                    @[simp]
                    theorem AddSubmonoid.mem_toSubsemigroup {M : Type u_1} [AddZeroClass M] {s : AddSubmonoid M} {x : M} :
                    x s.toAddSubsemigroup x s
                    @[simp]
                    theorem Submonoid.mem_toSubsemigroup {M : Type u_1} [MulOneClass M] {s : Submonoid M} {x : M} :
                    x s.toSubsemigroup x s
                    theorem AddSubmonoid.mem_carrier {M : Type u_1} [AddZeroClass M] {s : AddSubmonoid M} {x : M} :
                    x s.carrier x s
                    theorem Submonoid.mem_carrier {M : Type u_1} [MulOneClass M] {s : Submonoid M} {x : M} :
                    x s.carrier x s
                    @[simp]
                    theorem AddSubmonoid.mem_mk {M : Type u_1} [AddZeroClass M] {s : AddSubsemigroup M} {x : M} (h_one : 0 s.carrier) :
                    x { toAddSubsemigroup := s, zero_mem' := h_one } x s
                    @[simp]
                    theorem Submonoid.mem_mk {M : Type u_1} [MulOneClass M] {s : Subsemigroup M} {x : M} (h_one : 1 s.carrier) :
                    x { toSubsemigroup := s, one_mem' := h_one } x s
                    @[simp]
                    theorem AddSubmonoid.coe_set_mk {M : Type u_1} [AddZeroClass M] {s : AddSubsemigroup M} (h_one : 0 s.carrier) :
                    { toAddSubsemigroup := s, zero_mem' := h_one } = s
                    @[simp]
                    theorem Submonoid.coe_set_mk {M : Type u_1} [MulOneClass M] {s : Subsemigroup M} (h_one : 1 s.carrier) :
                    { toSubsemigroup := s, one_mem' := h_one } = s
                    @[simp]
                    theorem AddSubmonoid.mk_le_mk {M : Type u_1} [AddZeroClass M] {s : AddSubsemigroup M} {t : AddSubsemigroup M} (h_one : 0 s.carrier) (h_one' : 0 t.carrier) :
                    { toAddSubsemigroup := s, zero_mem' := h_one } { toAddSubsemigroup := t, zero_mem' := h_one' } s t
                    @[simp]
                    theorem Submonoid.mk_le_mk {M : Type u_1} [MulOneClass M] {s : Subsemigroup M} {t : Subsemigroup M} (h_one : 1 s.carrier) (h_one' : 1 t.carrier) :
                    { toSubsemigroup := s, one_mem' := h_one } { toSubsemigroup := t, one_mem' := h_one' } s t
                    theorem AddSubmonoid.ext {M : Type u_1} [AddZeroClass M] {S : AddSubmonoid M} {T : AddSubmonoid M} (h : ∀ (x : M), x S x T) :
                    S = T

                    Two AddSubmonoids are equal if they have the same elements.

                    theorem Submonoid.ext {M : Type u_1} [MulOneClass M] {S : Submonoid M} {T : Submonoid M} (h : ∀ (x : M), x S x T) :
                    S = T

                    Two submonoids are equal if they have the same elements.

                    def AddSubmonoid.copy {M : Type u_1} [AddZeroClass M] (S : AddSubmonoid M) (s : Set M) (hs : s = S) :

                    Copy an additive submonoid replacing carrier with a set that is equal to it.

                    Equations
                    • AddSubmonoid.copy S s hs = { toAddSubsemigroup := { carrier := s, add_mem' := (_ : ∀ {a b : M}, a sb sa + b s) }, zero_mem' := (_ : 0 s) }
                    Instances For
                      theorem AddSubmonoid.copy.proof_1 {M : Type u_1} [AddZeroClass M] (S : AddSubmonoid M) (s : Set M) (hs : s = S) :
                      ∀ {a b : M}, a sb sa + b s
                      theorem AddSubmonoid.copy.proof_2 {M : Type u_1} [AddZeroClass M] (S : AddSubmonoid M) (s : Set M) (hs : s = S) :
                      0 s
                      def Submonoid.copy {M : Type u_1} [MulOneClass M] (S : Submonoid M) (s : Set M) (hs : s = S) :

                      Copy a submonoid replacing carrier with a set that is equal to it.

                      Equations
                      • Submonoid.copy S s hs = { toSubsemigroup := { carrier := s, mul_mem' := (_ : ∀ {a b : M}, a sb sa * b s) }, one_mem' := (_ : 1 s) }
                      Instances For
                        @[simp]
                        theorem AddSubmonoid.coe_copy {M : Type u_1} [AddZeroClass M] {S : AddSubmonoid M} {s : Set M} (hs : s = S) :
                        ↑(AddSubmonoid.copy S s hs) = s
                        @[simp]
                        theorem Submonoid.coe_copy {M : Type u_1} [MulOneClass M] {S : Submonoid M} {s : Set M} (hs : s = S) :
                        ↑(Submonoid.copy S s hs) = s
                        theorem AddSubmonoid.copy_eq {M : Type u_1} [AddZeroClass M] {S : AddSubmonoid M} {s : Set M} (hs : s = S) :
                        theorem Submonoid.copy_eq {M : Type u_1} [MulOneClass M] {S : Submonoid M} {s : Set M} (hs : s = S) :
                        Submonoid.copy S s hs = S
                        theorem AddSubmonoid.zero_mem {M : Type u_1} [AddZeroClass M] (S : AddSubmonoid M) :
                        0 S

                        An AddSubmonoid contains the monoid's 0.

                        theorem Submonoid.one_mem {M : Type u_1} [MulOneClass M] (S : Submonoid M) :
                        1 S

                        A submonoid contains the monoid's 1.

                        theorem AddSubmonoid.add_mem {M : Type u_1} [AddZeroClass M] (S : AddSubmonoid M) {x : M} {y : M} :
                        x Sy Sx + y S

                        An AddSubmonoid is closed under addition.

                        theorem Submonoid.mul_mem {M : Type u_1} [MulOneClass M] (S : Submonoid M) {x : M} {y : M} :
                        x Sy Sx * y S

                        A submonoid is closed under multiplication.

                        theorem AddSubmonoid.instTopAddSubmonoid.proof_1 {M : Type u_1} [AddZeroClass M] :
                        ∀ {a b : M}, a Set.univb Set.univa + b Set.univ

                        The additive submonoid M of the AddMonoid M.

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

                        The submonoid M of the monoid M.

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

                        The trivial AddSubmonoid {0} of an AddMonoid M.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        theorem AddSubmonoid.instBotAddSubmonoid.proof_1 {M : Type u_1} [AddZeroClass M] :
                        ∀ {a b : M}, a {0}b {0}a + b {0}

                        The trivial submonoid {1} of a monoid M.

                        Equations
                        • Submonoid.instBotSubmonoid = { bot := { toSubsemigroup := { carrier := {1}, mul_mem' := (_ : ∀ {a b : M}, a {1}b {1}a * b {1}) }, one_mem' := (_ : 1 {1}) } }
                        Equations
                        • AddSubmonoid.instInhabitedAddSubmonoid = { default := }
                        Equations
                        • Submonoid.instInhabitedSubmonoid = { default := }
                        @[simp]
                        theorem AddSubmonoid.mem_bot {M : Type u_1} [AddZeroClass M] {x : M} :
                        x x = 0
                        @[simp]
                        theorem Submonoid.mem_bot {M : Type u_1} [MulOneClass M] {x : M} :
                        x x = 1
                        @[simp]
                        theorem AddSubmonoid.mem_top {M : Type u_1} [AddZeroClass M] (x : M) :
                        @[simp]
                        theorem Submonoid.mem_top {M : Type u_1} [MulOneClass M] (x : M) :
                        @[simp]
                        theorem AddSubmonoid.coe_top {M : Type u_1} [AddZeroClass M] :
                        = Set.univ
                        @[simp]
                        theorem Submonoid.coe_top {M : Type u_1} [MulOneClass M] :
                        = Set.univ
                        @[simp]
                        theorem AddSubmonoid.coe_bot {M : Type u_1} [AddZeroClass M] :
                        = {0}
                        @[simp]
                        theorem Submonoid.coe_bot {M : Type u_1} [MulOneClass M] :
                        = {1}
                        theorem AddSubmonoid.instInfAddSubmonoid.proof_1 {M : Type u_1} [AddZeroClass M] (S₁ : AddSubmonoid M) (S₂ : AddSubmonoid M) :
                        ∀ {a b : M}, a S₁ S₂b S₁ S₂a + b S₁ S₂
                        abbrev AddSubmonoid.instInfAddSubmonoid.match_1 {M : Type u_1} [AddZeroClass M] (S₁ : AddSubmonoid M) (S₂ : AddSubmonoid M) :
                        {b : M} → ∀ (motive : b S₁ S₂Prop) (x : b S₁ S₂), (∀ (hy : b S₁) (hy' : b S₂), motive (_ : b S₁ b S₂)) → motive x
                        Equations
                        Instances For
                          theorem AddSubmonoid.instInfAddSubmonoid.proof_2 {M : Type u_1} [AddZeroClass M] (S₁ : AddSubmonoid M) (S₂ : AddSubmonoid M) :
                          0 S₁ 0 S₂

                          The inf of two AddSubmonoids is their intersection.

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

                          The inf of two submonoids is their intersection.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          @[simp]
                          theorem AddSubmonoid.coe_inf {M : Type u_1} [AddZeroClass M] (p : AddSubmonoid M) (p' : AddSubmonoid M) :
                          ↑(p p') = p p'
                          @[simp]
                          theorem Submonoid.coe_inf {M : Type u_1} [MulOneClass M] (p : Submonoid M) (p' : Submonoid M) :
                          ↑(p p') = p p'
                          @[simp]
                          theorem AddSubmonoid.mem_inf {M : Type u_1} [AddZeroClass M] {p : AddSubmonoid M} {p' : AddSubmonoid M} {x : M} :
                          x p p' x p x p'
                          @[simp]
                          theorem Submonoid.mem_inf {M : Type u_1} [MulOneClass M] {p : Submonoid M} {p' : Submonoid M} {x : M} :
                          x p p' x p x p'
                          Equations
                          • One or more equations did not get rendered due to their size.
                          theorem AddSubmonoid.instInfSetAddSubmonoid.proof_1 {M : Type u_1} [AddZeroClass M] (s : Set (AddSubmonoid M)) :
                          ∀ {a b : M}, a ⋂ (t : AddSubmonoid M) (_ : t s), tb ⋂ (t : AddSubmonoid M) (_ : t s), ta + b ⋂ (x : AddSubmonoid M) (_ : x s), x
                          theorem AddSubmonoid.instInfSetAddSubmonoid.proof_2 {M : Type u_1} [AddZeroClass M] (s : Set (AddSubmonoid M)) :
                          0 ⋂ (x : AddSubmonoid M) (_ : x s), x
                          Equations
                          • One or more equations did not get rendered due to their size.
                          @[simp]
                          theorem AddSubmonoid.coe_sInf {M : Type u_1} [AddZeroClass M] (S : Set (AddSubmonoid M)) :
                          ↑(sInf S) = ⋂ (s : AddSubmonoid M) (_ : s S), s
                          @[simp]
                          theorem Submonoid.coe_sInf {M : Type u_1} [MulOneClass M] (S : Set (Submonoid M)) :
                          ↑(sInf S) = ⋂ (s : Submonoid M) (_ : s S), s
                          theorem AddSubmonoid.mem_sInf {M : Type u_1} [AddZeroClass M] {S : Set (AddSubmonoid M)} {x : M} :
                          x sInf S ∀ (p : AddSubmonoid M), p Sx p
                          theorem Submonoid.mem_sInf {M : Type u_1} [MulOneClass M] {S : Set (Submonoid M)} {x : M} :
                          x sInf S ∀ (p : Submonoid M), p Sx p
                          theorem AddSubmonoid.mem_iInf {M : Type u_1} [AddZeroClass M] {ι : Sort u_4} {S : ιAddSubmonoid M} {x : M} :
                          x ⨅ (i : ι), S i ∀ (i : ι), x S i
                          theorem Submonoid.mem_iInf {M : Type u_1} [MulOneClass M] {ι : Sort u_4} {S : ιSubmonoid M} {x : M} :
                          x ⨅ (i : ι), S i ∀ (i : ι), x S i
                          @[simp]
                          theorem AddSubmonoid.coe_iInf {M : Type u_1} [AddZeroClass M] {ι : Sort u_4} {S : ιAddSubmonoid M} :
                          ↑(⨅ (i : ι), S i) = ⋂ (i : ι), ↑(S i)
                          @[simp]
                          theorem Submonoid.coe_iInf {M : Type u_1} [MulOneClass M] {ι : Sort u_4} {S : ιSubmonoid M} :
                          ↑(⨅ (i : ι), S i) = ⋂ (i : ι), ↑(S i)
                          theorem AddSubmonoid.instCompleteLatticeAddSubmonoid.proof_9 {M : Type u_1} [AddZeroClass M] :
                          ∀ (x x_1 : AddSubmonoid M) (x_2 : M), x_2 x x_2 x_1x_2 x
                          theorem AddSubmonoid.instCompleteLatticeAddSubmonoid.proof_10 {M : Type u_1} [AddZeroClass M] :
                          ∀ (x x_1 : AddSubmonoid M) (x_2 : M), x_2 x x_2 x_1x_2 x_1
                          theorem AddSubmonoid.instCompleteLatticeAddSubmonoid.proof_13 {M : Type u_1} [AddZeroClass M] (s : Set (AddSubmonoid M)) (a : AddSubmonoid M) :
                          (∀ (b : AddSubmonoid M), b sb a) → sSup s a

                          The AddSubmonoids of an AddMonoid form a complete lattice.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          theorem AddSubmonoid.instCompleteLatticeAddSubmonoid.proof_15 {M : Type u_1} [AddZeroClass M] (s : Set (AddSubmonoid M)) (a : AddSubmonoid M) :
                          (∀ (b : AddSubmonoid M), b sa b) → a sInf s
                          theorem AddSubmonoid.instCompleteLatticeAddSubmonoid.proof_11 {M : Type u_1} [AddZeroClass M] :
                          ∀ (x x_1 x_2 : AddSubmonoid M), x x_1x x_2∀ (x_3 : M), x_3 xx_3 x_1 x_3 x_2

                          Submonoids of a monoid form a complete lattice.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Equations
                          • AddSubmonoid.instUniqueAddSubmonoid = { toInhabited := { default := }, uniq := (_ : ∀ (a : AddSubmonoid M), a = default) }
                          Equations
                          • Submonoid.instUniqueSubmonoid = { toInhabited := { default := }, uniq := (_ : ∀ (a : Submonoid M), a = default) }

                          The add_submonoid generated by a set

                          Equations
                          Instances For
                            def Submonoid.closure {M : Type u_1} [MulOneClass M] (s : Set M) :

                            The Submonoid generated by a set.

                            Equations
                            Instances For
                              theorem AddSubmonoid.mem_closure {M : Type u_1} [AddZeroClass M] {s : Set M} {x : M} :
                              x AddSubmonoid.closure s ∀ (S : AddSubmonoid M), s Sx S
                              theorem Submonoid.mem_closure {M : Type u_1} [MulOneClass M] {s : Set M} {x : M} :
                              x Submonoid.closure s ∀ (S : Submonoid M), s Sx S
                              @[simp]

                              The AddSubmonoid generated by a set includes the set.

                              @[simp]
                              theorem Submonoid.subset_closure {M : Type u_1} [MulOneClass M] {s : Set M} :

                              The submonoid generated by a set includes the set.

                              theorem Submonoid.not_mem_of_not_mem_closure {M : Type u_1} [MulOneClass M] {s : Set M} {P : M} (hP : ¬P Submonoid.closure s) :
                              ¬P s
                              @[simp]
                              theorem AddSubmonoid.closure_le {M : Type u_1} [AddZeroClass M] {s : Set M} {S : AddSubmonoid M} :

                              An additive submonoid S includes closure s if and only if it includes s

                              @[simp]
                              theorem Submonoid.closure_le {M : Type u_1} [MulOneClass M] {s : Set M} {S : Submonoid M} :

                              A submonoid S includes closure s if and only if it includes s.

                              theorem AddSubmonoid.closure_mono {M : Type u_1} [AddZeroClass M] ⦃s : Set M ⦃t : Set M (h : s t) :

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

                              theorem Submonoid.closure_mono {M : Type u_1} [MulOneClass M] ⦃s : Set M ⦃t : Set M (h : s t) :

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

                              theorem AddSubmonoid.closure_eq_of_le {M : Type u_1} [AddZeroClass M] {s : Set M} {S : AddSubmonoid M} (h₁ : s S) (h₂ : S AddSubmonoid.closure s) :
                              theorem Submonoid.closure_eq_of_le {M : Type u_1} [MulOneClass M] {s : Set M} {S : Submonoid M} (h₁ : s S) (h₂ : S Submonoid.closure s) :
                              theorem AddSubmonoid.closure_induction {M : Type u_1} [AddZeroClass M] {s : Set M} {p : MProp} {x : M} (h : x AddSubmonoid.closure s) (Hs : (x : M) → x sp x) (H1 : p 0) (Hmul : (x y : M) → p xp yp (x + y)) :
                              p x

                              An induction principle for additive closure membership. If p holds for 0 and all elements of s, and is preserved under addition, then p holds for all elements of the additive closure of s.

                              theorem Submonoid.closure_induction {M : Type u_1} [MulOneClass M] {s : Set M} {p : MProp} {x : M} (h : x Submonoid.closure s) (Hs : (x : M) → x sp x) (H1 : p 1) (Hmul : (x y : M) → p xp yp (x * y)) :
                              p x

                              An induction principle for closure membership. If p holds for 1 and all elements of s, and is preserved under multiplication, then p holds for all elements of the closure of s.

                              theorem AddSubmonoid.closure_induction' {M : Type u_1} [AddZeroClass M] (s : Set M) {p : (x : M) → x AddSubmonoid.closure sProp} (Hs : (x : M) → (h : x s) → p x (_ : x ↑(AddSubmonoid.closure s))) (H1 : p 0 (_ : 0 AddSubmonoid.closure s)) (Hmul : (x : M) → (hx : x AddSubmonoid.closure s) → (y : M) → (hy : y AddSubmonoid.closure s) → p x hxp y hyp (x + y) (_ : x + y AddSubmonoid.closure s)) {x : M} (hx : x AddSubmonoid.closure s) :
                              p x hx

                              A dependent version of AddSubmonoid.closure_induction.

                              abbrev AddSubmonoid.closure_induction'.match_1 {M : Type u_1} [AddZeroClass M] (s : Set M) {p : (x : M) → x AddSubmonoid.closure sProp} (y : M) (motive : (x, p y x) → Prop) :
                              (x : x, p y x) → ((hy' : y AddSubmonoid.closure s) → (hy : p y hy') → motive (_ : x, p y x)) → motive x
                              Equations
                              Instances For
                                theorem Submonoid.closure_induction' {M : Type u_1} [MulOneClass M] (s : Set M) {p : (x : M) → x Submonoid.closure sProp} (Hs : (x : M) → (h : x s) → p x (_ : x ↑(Submonoid.closure s))) (H1 : p 1 (_ : 1 Submonoid.closure s)) (Hmul : (x : M) → (hx : x Submonoid.closure s) → (y : M) → (hy : y Submonoid.closure s) → p x hxp y hyp (x * y) (_ : x * y Submonoid.closure s)) {x : M} (hx : x Submonoid.closure s) :
                                p x hx

                                A dependent version of Submonoid.closure_induction.

                                theorem AddSubmonoid.closure_induction₂ {M : Type u_1} [AddZeroClass M] {s : Set M} {p : MMProp} {x : M} {y : M} (hx : x AddSubmonoid.closure s) (hy : y AddSubmonoid.closure s) (Hs : (x : M) → x s(y : M) → y sp x y) (H1_left : (x : M) → p 0 x) (H1_right : (x : M) → p x 0) (Hmul_left : (x y z : M) → p x zp y zp (x + y) z) (Hmul_right : (x y z : M) → p z xp z yp z (x + y)) :
                                p x y

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

                                theorem Submonoid.closure_induction₂ {M : Type u_1} [MulOneClass M] {s : Set M} {p : MMProp} {x : M} {y : M} (hx : x Submonoid.closure s) (hy : y Submonoid.closure s) (Hs : (x : M) → x s(y : M) → y sp x y) (H1_left : (x : M) → p 1 x) (H1_right : (x : M) → p x 1) (Hmul_left : (x y z : M) → p x zp y zp (x * y) z) (Hmul_right : (x y z : M) → p z xp z yp z (x * y)) :
                                p x y

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

                                theorem AddSubmonoid.dense_induction {M : Type u_1} [AddZeroClass M] {p : MProp} (x : M) {s : Set M} (hs : AddSubmonoid.closure s = ) (Hs : (x : M) → x sp x) (H1 : p 0) (Hmul : (x y : M) → p xp yp (x + y)) :
                                p x

                                If s is a dense set in an additive monoid M, AddSubmonoid.closure s = ⊤, then in order to prove that some predicate p holds for all x : M it suffices to verify p x for x ∈ s, verify p 0, and verify that p x and p y imply p (x + y).

                                theorem Submonoid.dense_induction {M : Type u_1} [MulOneClass M] {p : MProp} (x : M) {s : Set M} (hs : Submonoid.closure s = ) (Hs : (x : M) → x sp x) (H1 : p 1) (Hmul : (x y : M) → p xp yp (x * y)) :
                                p x

                                If s is a dense set in a monoid M, Submonoid.closure s = ⊤, then in order to prove that some predicate p holds for all x : M it suffices to verify p x for x ∈ s, verify p 1, and verify that p x and p y imply p (x * y).

                                theorem AddSubmonoid.gi.proof_1 (M : Type u_1) [AddZeroClass M] :
                                ∀ (x : AddSubmonoid M), x ↑(AddSubmonoid.closure x)
                                theorem AddSubmonoid.gi.proof_2 (M : Type u_1) [AddZeroClass M] :
                                ∀ (x : Set M) (x_1 : ↑(AddSubmonoid.closure x) x), (fun s x => AddSubmonoid.closure s) x x_1 = (fun s x => AddSubmonoid.closure s) x x_1
                                def AddSubmonoid.gi (M : Type u_1) [AddZeroClass M] :
                                GaloisInsertion AddSubmonoid.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
                                  def Submonoid.gi (M : Type u_1) [MulOneClass M] :
                                  GaloisInsertion Submonoid.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
                                    @[simp]

                                    Additive closure of an additive submonoid S equals S

                                    @[simp]
                                    theorem Submonoid.closure_eq {M : Type u_1} [MulOneClass M] (S : Submonoid M) :

                                    Closure of a submonoid S equals S.

                                    @[simp]
                                    theorem Submonoid.sup_eq_closure {M : Type u_1} [MulOneClass M] (N : Submonoid M) (N' : Submonoid M) :
                                    N N' = Submonoid.closure (N N')
                                    theorem AddSubmonoid.closure_iUnion {M : Type u_1} [AddZeroClass M] {ι : Sort u_4} (s : ιSet M) :
                                    AddSubmonoid.closure (⋃ (i : ι), s i) = ⨆ (i : ι), AddSubmonoid.closure (s i)
                                    theorem Submonoid.closure_iUnion {M : Type u_1} [MulOneClass M] {ι : Sort u_4} (s : ιSet M) :
                                    Submonoid.closure (⋃ (i : ι), s i) = ⨆ (i : ι), Submonoid.closure (s i)
                                    theorem AddSubmonoid.mem_iSup {M : Type u_1} [AddZeroClass M] {ι : Sort u_4} (p : ιAddSubmonoid M) {m : M} :
                                    m ⨆ (i : ι), p i ∀ (N : AddSubmonoid M), (∀ (i : ι), p i N) → m N
                                    theorem Submonoid.mem_iSup {M : Type u_1} [MulOneClass M] {ι : Sort u_4} (p : ιSubmonoid M) {m : M} :
                                    m ⨆ (i : ι), p i ∀ (N : Submonoid M), (∀ (i : ι), p i N) → m N
                                    theorem AddSubmonoid.iSup_eq_closure {M : Type u_1} [AddZeroClass M] {ι : Sort u_4} (p : ιAddSubmonoid M) :
                                    ⨆ (i : ι), p i = AddSubmonoid.closure (⋃ (i : ι), ↑(p i))
                                    theorem Submonoid.iSup_eq_closure {M : Type u_1} [MulOneClass M] {ι : Sort u_4} (p : ιSubmonoid M) :
                                    ⨆ (i : ι), p i = Submonoid.closure (⋃ (i : ι), ↑(p i))
                                    theorem AddSubmonoid.disjoint_def {M : Type u_1} [AddZeroClass M] {p₁ : AddSubmonoid M} {p₂ : AddSubmonoid M} :
                                    Disjoint p₁ p₂ ∀ {x : M}, x p₁x p₂x = 0
                                    theorem Submonoid.disjoint_def {M : Type u_1} [MulOneClass M] {p₁ : Submonoid M} {p₂ : Submonoid M} :
                                    Disjoint p₁ p₂ ∀ {x : M}, x p₁x p₂x = 1
                                    theorem AddSubmonoid.disjoint_def' {M : Type u_1} [AddZeroClass M] {p₁ : AddSubmonoid M} {p₂ : AddSubmonoid M} :
                                    Disjoint p₁ p₂ ∀ {x y : M}, x p₁y p₂x = yx = 0
                                    theorem Submonoid.disjoint_def' {M : Type u_1} [MulOneClass M] {p₁ : Submonoid M} {p₂ : Submonoid M} :
                                    Disjoint p₁ p₂ ∀ {x y : M}, x p₁y p₂x = yx = 1
                                    theorem AddMonoidHom.eqLocusM.proof_1 {M : Type u_2} {N : Type u_1} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (g : M →+ N) :
                                    ∀ {a b : M}, f a = g af b = g bf (a + b) = g (a + b)
                                    theorem AddMonoidHom.eqLocusM.proof_2 {M : Type u_1} {N : Type u_2} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (g : M →+ N) :
                                    0 { carrier := {x | f x = g x}, add_mem' := (_ : ∀ {a b : M}, f a = g af b = g bf (a + b) = g (a + b)) }.carrier
                                    def AddMonoidHom.eqLocusM {M : Type u_1} {N : Type u_2} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (g : M →+ N) :

                                    The additive submonoid of elements x : M such that f x = g x

                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    Instances For
                                      def MonoidHom.eqLocusM {M : Type u_1} {N : Type u_2} [MulOneClass M] [MulOneClass N] (f : M →* N) (g : M →* N) :

                                      The submonoid of elements x : M such that f x = g x

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For
                                        @[simp]
                                        theorem AddMonoidHom.eqLocusM_same {M : Type u_1} {N : Type u_2} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) :
                                        @[simp]
                                        theorem MonoidHom.eqLocusM_same {M : Type u_1} {N : Type u_2} [MulOneClass M] [MulOneClass N] (f : M →* N) :
                                        theorem AddMonoidHom.eqOn_closureM {M : Type u_1} {N : Type u_2} [AddZeroClass M] [AddZeroClass N] {f : M →+ N} {g : M →+ N} {s : Set M} (h : Set.EqOn (f) (g) s) :

                                        If two monoid homomorphisms are equal on a set, then they are equal on its submonoid closure.

                                        theorem MonoidHom.eqOn_closureM {M : Type u_1} {N : Type u_2} [MulOneClass M] [MulOneClass N] {f : M →* N} {g : M →* N} {s : Set M} (h : Set.EqOn (f) (g) s) :
                                        Set.EqOn f g ↑(Submonoid.closure s)

                                        If two monoid homomorphisms are equal on a set, then they are equal on its submonoid closure.

                                        theorem AddMonoidHom.eq_of_eqOn_topM {M : Type u_1} {N : Type u_2} [AddZeroClass M] [AddZeroClass N] {f : M →+ N} {g : M →+ N} (h : Set.EqOn f g ) :
                                        f = g
                                        theorem MonoidHom.eq_of_eqOn_topM {M : Type u_1} {N : Type u_2} [MulOneClass M] [MulOneClass N] {f : M →* N} {g : M →* N} (h : Set.EqOn f g ) :
                                        f = g
                                        theorem AddMonoidHom.eq_of_eqOn_denseM {M : Type u_1} {N : Type u_2} [AddZeroClass M] [AddZeroClass N] {s : Set M} (hs : AddSubmonoid.closure s = ) {f : M →+ N} {g : M →+ N} (h : Set.EqOn (f) (g) s) :
                                        f = g
                                        theorem MonoidHom.eq_of_eqOn_denseM {M : Type u_1} {N : Type u_2} [MulOneClass M] [MulOneClass N] {s : Set M} (hs : Submonoid.closure s = ) {f : M →* N} {g : M →* N} (h : Set.EqOn (f) (g) s) :
                                        f = g

                                        The additive submonoid consisting of the additive units of an additive monoid

                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        Instances For
                                          theorem IsAddUnit.addSubmonoid.proof_1 (M : Type u_1) [AddMonoid M] {a : M} {b : M} (ha : a setOf IsAddUnit) (hb : b setOf IsAddUnit) :
                                          a + b setOf IsAddUnit
                                          def IsUnit.submonoid (M : Type u_4) [Monoid M] :

                                          The submonoid consisting of the units of a monoid

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            theorem AddMonoidHom.ofClosureMEqTopLeft.proof_1 {M : Type u_2} {N : Type u_1} [AddMonoid M] [AddMonoid N] {s : Set M} (f : MN) (hs : AddSubmonoid.closure s = ) (h1 : f 0 = 0) (hmul : ∀ (x : M), x s∀ (y : M), f (x + y) = f x + f y) (x : M) (y : M) :
                                            f (x + y) = f x + f y
                                            def AddMonoidHom.ofClosureMEqTopLeft {M : Type u_4} {N : Type u_5} [AddMonoid M] [AddMonoid N] {s : Set M} (f : MN) (hs : AddSubmonoid.closure s = ) (h1 : f 0 = 0) (hmul : ∀ (x : M), x s∀ (y : M), f (x + y) = f x + f y) :
                                            M →+ N

                                            Let s be a subset of an additive monoid M such that the closure of s is the whole monoid. Then AddMonoidHom.ofClosureEqTopLeft defines an additive monoid homomorphism from M asking for a proof of f (x + y) = f x + f y only for x ∈ s.

                                            Equations
                                            Instances For
                                              def MonoidHom.ofClosureMEqTopLeft {M : Type u_4} {N : Type u_5} [Monoid M] [Monoid N] {s : Set M} (f : MN) (hs : Submonoid.closure s = ) (h1 : f 1 = 1) (hmul : ∀ (x : M), x s∀ (y : M), f (x * y) = f x * f y) :
                                              M →* N

                                              Let s be a subset of a monoid M such that the closure of s is the whole monoid. Then MonoidHom.ofClosureEqTopLeft defines a monoid homomorphism from M asking for a proof of f (x * y) = f x * f y only for x ∈ s.

                                              Equations
                                              Instances For
                                                @[simp]
                                                theorem AddMonoidHom.coe_ofClosureMEqTopLeft {M : Type u_1} {N : Type u_2} [AddMonoid M] [AddMonoid N] {s : Set M} (f : MN) (hs : AddSubmonoid.closure s = ) (h1 : f 0 = 0) (hmul : ∀ (x : M), x s∀ (y : M), f (x + y) = f x + f y) :
                                                @[simp]
                                                theorem MonoidHom.coe_ofClosureMEqTopLeft {M : Type u_1} {N : Type u_2} [Monoid M] [Monoid N] {s : Set M} (f : MN) (hs : Submonoid.closure s = ) (h1 : f 1 = 1) (hmul : ∀ (x : M), x s∀ (y : M), f (x * y) = f x * f y) :
                                                ↑(MonoidHom.ofClosureMEqTopLeft f hs h1 hmul) = f
                                                def AddMonoidHom.ofClosureMEqTopRight {M : Type u_4} {N : Type u_5} [AddMonoid M] [AddMonoid N] {s : Set M} (f : MN) (hs : AddSubmonoid.closure s = ) (h1 : f 0 = 0) (hmul : ∀ (x y : M), y sf (x + y) = f x + f y) :
                                                M →+ N

                                                Let s be a subset of an additive monoid M such that the closure of s is the whole monoid. Then AddMonoidHom.ofClosureEqTopRight defines an additive monoid homomorphism from M asking for a proof of f (x + y) = f x + f y only for y ∈ s.

                                                Equations
                                                • One or more equations did not get rendered due to their size.
                                                Instances For
                                                  theorem AddMonoidHom.ofClosureMEqTopRight.proof_1 {M : Type u_2} {N : Type u_1} [AddMonoid M] [AddMonoid N] {s : Set M} (f : MN) (hs : AddSubmonoid.closure s = ) (h1 : f 0 = 0) (hmul : ∀ (x y : M), y sf (x + y) = f x + f y) (x : M) (y : M) :
                                                  ZeroHom.toFun { toFun := f, map_zero' := h1 } (x + y) = ZeroHom.toFun { toFun := f, map_zero' := h1 } x + ZeroHom.toFun { toFun := f, map_zero' := h1 } y
                                                  def MonoidHom.ofClosureMEqTopRight {M : Type u_4} {N : Type u_5} [Monoid M] [Monoid N] {s : Set M} (f : MN) (hs : Submonoid.closure s = ) (h1 : f 1 = 1) (hmul : ∀ (x y : M), y sf (x * y) = f x * f y) :
                                                  M →* N

                                                  Let s be a subset of a monoid M such that the closure of s is the whole monoid. Then MonoidHom.ofClosureEqTopRight defines a monoid homomorphism from M asking for a proof of f (x * y) = f x * f y only for y ∈ s.

                                                  Equations
                                                  • One or more equations did not get rendered due to their size.
                                                  Instances For
                                                    @[simp]
                                                    theorem AddMonoidHom.coe_ofClosureMEqTopRight {M : Type u_1} {N : Type u_2} [AddMonoid M] [AddMonoid N] {s : Set M} (f : MN) (hs : AddSubmonoid.closure s = ) (h1 : f 0 = 0) (hmul : ∀ (x y : M), y sf (x + y) = f x + f y) :
                                                    @[simp]
                                                    theorem MonoidHom.coe_ofClosureMEqTopRight {M : Type u_1} {N : Type u_2} [Monoid M] [Monoid N] {s : Set M} (f : MN) (hs : Submonoid.closure s = ) (h1 : f 1 = 1) (hmul : ∀ (x y : M), y sf (x * y) = f x * f y) :
                                                    ↑(MonoidHom.ofClosureMEqTopRight f hs h1 hmul) = f