Documentation

Mathlib.Algebra.Group.Subsemigroup.Defs

Subsemigroups: definition #

This file defines bundled multiplicative and additive subsemigroups.

Main definitions #

For each of the following definitions in the Subsemigroup namespace, there is a corresponding definition in the AddSubsemigroup namespace.

Similarly, for each of these definitions in the MulHom namespace, there is a corresponding definition in the AddHom namespace.

Implementation notes #

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

Note that Subsemigroup M does not actually require Semigroup M, instead requiring only the weaker Mul M.

This file is designed to have very few dependencies. In particular, it should not use natural numbers.

Tags #

subsemigroup, subsemigroups

class MulMemClass (S : Type u_3) (M : outParam (Type u_4)) [Mul M] [SetLike S M] :

MulMemClass S M says S is a type of sets s : Set M that are closed under (*)

  • mul_mem {s : S} {a b : M} : a sb sa * b s

    A substructure satisfying MulMemClass is closed under multiplication.

Instances
    class AddMemClass (S : Type u_3) (M : outParam (Type u_4)) [Add M] [SetLike S M] :

    AddMemClass S M says S is a type of sets s : Set M that are closed under (+)

    • add_mem {s : S} {a b : M} : a sb sa + b s

      A substructure satisfying AddMemClass is closed under addition.

    Instances
      structure Subsemigroup (M : Type u_3) [Mul M] :
      Type u_3

      A subsemigroup of a magma M is a subset closed under multiplication.

      • carrier : Set M

        The carrier of a subsemigroup.

      • mul_mem' {a b : M} : a self.carrierb self.carriera * b self.carrier

        The product of two elements of a subsemigroup belongs to the subsemigroup.

      Instances For
        structure AddSubsemigroup (M : Type u_3) [Add M] :
        Type u_3

        An additive subsemigroup of an additive magma M is a subset closed under addition.

        • carrier : Set M

          The carrier of an additive subsemigroup.

        • add_mem' {a b : M} : a self.carrierb self.carriera + b self.carrier

          The sum of two elements of an additive subsemigroup belongs to the subsemigroup.

        Instances For
          instance Subsemigroup.instSetLike {M : Type u_1} [Mul M] :
          Equations
            Equations
              def Subsemigroup.ofClass {S : Type u_3} {M : Type u_4} [Mul M] [SetLike S M] [MulMemClass S M] (s : S) :

              The actual Subsemigroup obtained from an element of a MulMemClass.

              Equations
                Instances For
                  def AddSubsemigroup.ofClass {S : Type u_3} {M : Type u_4} [Add M] [SetLike S M] [AddMemClass S M] (s : S) :

                  The actual AddSubsemigroup obtained from an element of a AddMemClass

                  Equations
                    Instances For
                      @[simp]
                      theorem AddSubsemigroup.coe_ofClass {S : Type u_3} {M : Type u_4} [Add M] [SetLike S M] [AddMemClass S M] (s : S) :
                      (ofClass s) = s
                      @[simp]
                      theorem Subsemigroup.coe_ofClass {S : Type u_3} {M : Type u_4} [Mul M] [SetLike S M] [MulMemClass S M] (s : S) :
                      (ofClass s) = s
                      @[instance 100]
                      instance Subsemigroup.instCanLiftSetCoeForallForallForallMemForallHMul {M : Type u_1} [Mul M] :
                      CanLift (Set M) (Subsemigroup M) SetLike.coe fun (s : Set M) => ∀ {x y : M}, x sy sx * y s
                      @[instance 100]
                      instance AddSubsemigroup.instCanLiftSetCoeForallForallForallMemForallHAdd {M : Type u_1} [Add M] :
                      CanLift (Set M) (AddSubsemigroup M) SetLike.coe fun (s : Set M) => ∀ {x y : M}, x sy sx + y s
                      @[simp]
                      theorem Subsemigroup.mem_carrier {M : Type u_1} [Mul M] {s : Subsemigroup M} {x : M} :
                      x s.carrier x s
                      @[simp]
                      theorem AddSubsemigroup.mem_carrier {M : Type u_1} [Add M] {s : AddSubsemigroup M} {x : M} :
                      x s.carrier x s
                      @[simp]
                      theorem Subsemigroup.mem_mk {M : Type u_1} [Mul M] {s : Set M} {x : M} (h_mul : ∀ {a b : M}, a sb sa * b s) :
                      x { carrier := s, mul_mem' := h_mul } x s
                      @[simp]
                      theorem AddSubsemigroup.mem_mk {M : Type u_1} [Add M] {s : Set M} {x : M} (h_mul : ∀ {a b : M}, a sb sa + b s) :
                      x { carrier := s, add_mem' := h_mul } x s
                      @[simp]
                      theorem Subsemigroup.coe_set_mk {M : Type u_1} [Mul M] (s : Set M) (h_mul : ∀ {a b : M}, a sb sa * b s) :
                      { carrier := s, mul_mem' := h_mul } = s
                      @[simp]
                      theorem AddSubsemigroup.coe_set_mk {M : Type u_1} [Add M] (s : Set M) (h_mul : ∀ {a b : M}, a sb sa + b s) :
                      { carrier := s, add_mem' := h_mul } = s
                      @[simp]
                      theorem Subsemigroup.mk_le_mk {M : Type u_1} [Mul M] {s t : Set M} (h_mul : ∀ {a b : M}, a sb sa * b s) (h_mul' : ∀ {a b : M}, a tb ta * b t) :
                      { carrier := s, mul_mem' := h_mul } { carrier := t, mul_mem' := h_mul' } s t
                      @[simp]
                      theorem AddSubsemigroup.mk_le_mk {M : Type u_1} [Add M] {s t : Set M} (h_mul : ∀ {a b : M}, a sb sa + b s) (h_mul' : ∀ {a b : M}, a tb ta + b t) :
                      { carrier := s, add_mem' := h_mul } { carrier := t, add_mem' := h_mul' } s t
                      theorem Subsemigroup.ext {M : Type u_1} [Mul M] {S T : Subsemigroup M} (h : ∀ (x : M), x S x T) :
                      S = T

                      Two subsemigroups are equal if they have the same elements.

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

                      Two AddSubsemigroups are equal if they have the same elements.

                      theorem AddSubsemigroup.ext_iff {M : Type u_1} [Add M] {S T : AddSubsemigroup M} :
                      S = T ∀ (x : M), x S x T
                      theorem Subsemigroup.ext_iff {M : Type u_1} [Mul M] {S T : Subsemigroup M} :
                      S = T ∀ (x : M), x S x T
                      def Subsemigroup.copy {M : Type u_1} [Mul M] (S : Subsemigroup M) (s : Set M) (hs : s = S) :

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

                      Equations
                        Instances For
                          def AddSubsemigroup.copy {M : Type u_1} [Add M] (S : AddSubsemigroup M) (s : Set M) (hs : s = S) :

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

                          Equations
                            Instances For
                              @[simp]
                              theorem Subsemigroup.coe_copy {M : Type u_1} [Mul M] {S : Subsemigroup M} {s : Set M} (hs : s = S) :
                              (S.copy s hs) = s
                              @[simp]
                              theorem AddSubsemigroup.coe_copy {M : Type u_1} [Add M] {S : AddSubsemigroup M} {s : Set M} (hs : s = S) :
                              (S.copy s hs) = s
                              theorem Subsemigroup.copy_eq {M : Type u_1} [Mul M] {S : Subsemigroup M} {s : Set M} (hs : s = S) :
                              S.copy s hs = S
                              theorem AddSubsemigroup.copy_eq {M : Type u_1} [Add M] {S : AddSubsemigroup M} {s : Set M} (hs : s = S) :
                              S.copy s hs = S
                              theorem Subsemigroup.mul_mem {M : Type u_1} [Mul M] (S : Subsemigroup M) {x y : M} :
                              x Sy Sx * y S

                              A subsemigroup is closed under multiplication.

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

                              An AddSubsemigroup is closed under addition.

                              instance Subsemigroup.instTop {M : Type u_1} [Mul M] :

                              The subsemigroup M of the magma M.

                              Equations
                                instance AddSubsemigroup.instTop {M : Type u_1} [Add M] :

                                The additive subsemigroup M of the magma M.

                                Equations
                                  instance Subsemigroup.instBot {M : Type u_1} [Mul M] :

                                  The trivial subsemigroup of a magma M.

                                  Equations
                                    instance AddSubsemigroup.instBot {M : Type u_1} [Add M] :

                                    The trivial AddSubsemigroup of an additive magma M.

                                    Equations
                                      Equations
                                        theorem Subsemigroup.notMem_bot {M : Type u_1} [Mul M] {x : M} :
                                        x
                                        theorem AddSubsemigroup.notMem_bot {M : Type u_1} [Add M] {x : M} :
                                        x
                                        @[deprecated AddSubsemigroup.notMem_bot (since := "2025-05-23")]
                                        theorem AddSubsemigroup.not_mem_bot {M : Type u_1} [Add M] {x : M} :
                                        x

                                        Alias of AddSubsemigroup.notMem_bot.

                                        @[deprecated Subsemigroup.notMem_bot (since := "2025-05-23")]
                                        theorem Subsemigroup.not_mem_bot {M : Type u_1} [Mul M] {x : M} :
                                        x

                                        Alias of Subsemigroup.notMem_bot.

                                        @[simp]
                                        theorem Subsemigroup.mem_top {M : Type u_1} [Mul M] (x : M) :
                                        @[simp]
                                        theorem AddSubsemigroup.mem_top {M : Type u_1} [Add M] (x : M) :
                                        @[simp]
                                        theorem Subsemigroup.coe_top {M : Type u_1} [Mul M] :
                                        @[simp]
                                        theorem AddSubsemigroup.coe_top {M : Type u_1} [Add M] :
                                        @[simp]
                                        theorem Subsemigroup.coe_bot {M : Type u_1} [Mul M] :
                                        =
                                        @[simp]
                                        theorem AddSubsemigroup.coe_bot {M : Type u_1} [Add M] :
                                        =
                                        instance Subsemigroup.instMin {M : Type u_1} [Mul M] :

                                        The inf of two subsemigroups is their intersection.

                                        Equations
                                          instance AddSubsemigroup.instMin {M : Type u_1} [Add M] :

                                          The inf of two AddSubsemigroups is their intersection.

                                          Equations
                                            @[simp]
                                            theorem Subsemigroup.coe_inf {M : Type u_1} [Mul M] (p p' : Subsemigroup M) :
                                            (pp') = p p'
                                            @[simp]
                                            theorem AddSubsemigroup.coe_inf {M : Type u_1} [Add M] (p p' : AddSubsemigroup M) :
                                            (pp') = p p'
                                            @[simp]
                                            theorem Subsemigroup.mem_inf {M : Type u_1} [Mul M] {p p' : Subsemigroup M} {x : M} :
                                            x pp' x p x p'
                                            @[simp]
                                            theorem AddSubsemigroup.mem_inf {M : Type u_1} [Add M] {p p' : AddSubsemigroup M} {x : M} :
                                            x pp' x p x p'
                                            def MulHom.eqLocus {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] (f g : M →ₙ* N) :

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

                                            Equations
                                              Instances For
                                                def AddHom.eqLocus {M : Type u_1} {N : Type u_2} [Add M] [Add N] (f g : M →ₙ+ N) :

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

                                                Equations
                                                  Instances For
                                                    theorem MulHom.eq_of_eqOn_top {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] {f g : M →ₙ* N} (h : Set.EqOn f g ) :
                                                    f = g
                                                    theorem AddHom.eq_of_eqOn_top {M : Type u_1} {N : Type u_2} [Add M] [Add N] {f g : M →ₙ+ N} (h : Set.EqOn f g ) :
                                                    f = g
                                                    @[instance 900]
                                                    instance MulMemClass.mul {M : Type u_1} {A : Type u_3} [Mul M] [SetLike A M] [hA : MulMemClass A M] (S' : A) :
                                                    Mul S'

                                                    A submagma of a magma inherits a multiplication.

                                                    Equations
                                                      @[instance 900]
                                                      instance AddMemClass.add {M : Type u_1} {A : Type u_3} [Add M] [SetLike A M] [hA : AddMemClass A M] (S' : A) :
                                                      Add S'

                                                      An additive submagma of an additive magma inherits an addition.

                                                      Equations
                                                        @[simp]
                                                        theorem MulMemClass.coe_mul {M : Type u_1} {A : Type u_3} [Mul M] [SetLike A M] [hA : MulMemClass A M] (S' : A) (x y : S') :
                                                        ↑(x * y) = x * y
                                                        @[simp]
                                                        theorem AddMemClass.coe_add {M : Type u_1} {A : Type u_3} [Add M] [SetLike A M] [hA : AddMemClass A M] (S' : A) (x y : S') :
                                                        ↑(x + y) = x + y
                                                        @[simp]
                                                        theorem MulMemClass.mk_mul_mk {M : Type u_1} {A : Type u_3} [Mul M] [SetLike A M] [hA : MulMemClass A M] (S' : A) (x y : M) (hx : x S') (hy : y S') :
                                                        x, hx * y, hy = x * y,
                                                        @[simp]
                                                        theorem AddMemClass.mk_add_mk {M : Type u_1} {A : Type u_3} [Add M] [SetLike A M] [hA : AddMemClass A M] (S' : A) (x y : M) (hx : x S') (hy : y S') :
                                                        x, hx + y, hy = x + y,
                                                        theorem MulMemClass.mul_def {M : Type u_1} {A : Type u_3} [Mul M] [SetLike A M] [hA : MulMemClass A M] (S' : A) (x y : S') :
                                                        x * y = x * y,
                                                        theorem AddMemClass.add_def {M : Type u_1} {A : Type u_3} [Add M] [SetLike A M] [hA : AddMemClass A M] (S' : A) (x y : S') :
                                                        x + y = x + y,
                                                        instance MulMemClass.toSemigroup {M : Type u_4} [Semigroup M] {A : Type u_5} [SetLike A M] [MulMemClass A M] (S : A) :

                                                        A subsemigroup of a semigroup inherits a semigroup structure.

                                                        Equations
                                                          instance AddMemClass.toAddSemigroup {M : Type u_4} [AddSemigroup M] {A : Type u_5} [SetLike A M] [AddMemClass A M] (S : A) :

                                                          An AddSubsemigroup of an AddSemigroup inherits an AddSemigroup structure.

                                                          Equations
                                                            instance MulMemClass.toCommSemigroup {M : Type u_5} [CommSemigroup M] {A : Type u_4} [SetLike A M] [MulMemClass A M] (S : A) :

                                                            A subsemigroup of a CommSemigroup is a CommSemigroup.

                                                            Equations
                                                              def MulMemClass.subtype {M : Type u_1} {A : Type u_3} [Mul M] [SetLike A M] [hA : MulMemClass A M] (S' : A) :
                                                              S' →ₙ* M

                                                              The natural semigroup hom from a subsemigroup of semigroup M to M.

                                                              Equations
                                                                Instances For
                                                                  def AddMemClass.subtype {M : Type u_1} {A : Type u_3} [Add M] [SetLike A M] [hA : AddMemClass A M] (S' : A) :
                                                                  S' →ₙ+ M

                                                                  The natural semigroup hom from an AddSubsemigroup of AddSubsemigroup M to M.

                                                                  Equations
                                                                    Instances For
                                                                      @[simp]
                                                                      theorem MulMemClass.subtype_apply {M : Type u_1} {A : Type u_3} [Mul M] [SetLike A M] [hA : MulMemClass A M] {S' : A} (x : S') :
                                                                      (subtype S') x = x
                                                                      @[simp]
                                                                      theorem AddMemClass.subtype_apply {M : Type u_1} {A : Type u_3} [Add M] [SetLike A M] [hA : AddMemClass A M] {S' : A} (x : S') :
                                                                      (subtype S') x = x
                                                                      theorem MulMemClass.subtype_injective {M : Type u_1} {A : Type u_3} [Mul M] [SetLike A M] [hA : MulMemClass A M] (S' : A) :
                                                                      theorem AddMemClass.subtype_injective {M : Type u_1} {A : Type u_3} [Add M] [SetLike A M] [hA : AddMemClass A M] (S' : A) :
                                                                      @[simp]
                                                                      theorem MulMemClass.coe_subtype {M : Type u_1} {A : Type u_3} [Mul M] [SetLike A M] [hA : MulMemClass A M] (S' : A) :
                                                                      @[simp]
                                                                      theorem AddMemClass.coe_subtype {M : Type u_1} {A : Type u_3} [Add M] [SetLike A M] [hA : AddMemClass A M] (S' : A) :