Documentation

Mathlib.GroupTheory.Submonoid.Center

Centers of monoids #

Main definitions #

We provide Subgroup.center, AddSubgroup.center, Subsemiring.center, and Subring.center in other files.

The center of a multiplication with unit M is the set of elements that commute with everything in M

Equations
Instances For

    The center of an addition with zero M is the set of elements that commute with everything in M

    Equations
    Instances For
      @[reducible, inline]

      The center of a multiplication with unit is commutative and associative.

      This is not an instance as it forms an non-defeq diamond with Submonoid.toMonoid in the npow field.

      Equations
      Instances For
        @[reducible, inline]

        The center of an addition with zero is commutative and associative.

        Equations
        Instances For

          The center of a monoid is commutative.

          Equations
          theorem Submonoid.mem_center_iff {M : Type u_1} [Monoid M] {z : M} :
          z center M ∀ (g : M), g * z = z * g
          theorem AddSubmonoid.mem_center_iff {M : Type u_1} [AddMonoid M] {z : M} :
          z center M ∀ (g : M), g + z = z + g
          instance Submonoid.decidableMemCenter {M : Type u_1} [Monoid M] (a : M) [Decidable (∀ (b : M), b * a = a * b)] :
          Equations
          instance AddSubmonoid.decidableMemCenter {M : Type u_1} [AddMonoid M] (a : M) [Decidable (∀ (b : M), b + a = a + b)] :
          Equations

          The center of a monoid acts commutatively on that monoid.

          The center of a monoid acts commutatively on that monoid.

          Note that smulCommClass (center M) (center M) M is already implied by Submonoid.smulCommClass_right

          @[simp]

          For a monoid, the units of the center inject into the center of the units. This is not an equivalence in general; one case when it is is for groups with zero, which is covered in centerUnitsEquivUnitsCenter.

          Equations
          Instances For

            For an additive monoid, the units of the center inject into the center of the units.

            Equations
            Instances For
              @[simp]
              theorem val_unitsCenterToCenterUnits_apply_coe (M : Type u_1) [Monoid M] (n : (↥(Submonoid.center M))ˣ) :
              ((unitsCenterToCenterUnits M) n) = n
              theorem MulEquivClass.apply_mem_center {M : Type u_3} {N : Type u_1} {F : Type u_2} [EquivLike F M N] [Mul M] [Mul N] [MulEquivClass F M N] (e : F) {x : M} (hx : x Set.center M) :
              theorem AddEquivClass.apply_mem_center {M : Type u_3} {N : Type u_1} {F : Type u_2} [EquivLike F M N] [Add M] [Add N] [AddEquivClass F M N] (e : F) {x : M} (hx : x Set.addCenter M) :
              theorem MulEquivClass.apply_mem_center_iff {M : Type u_3} {N : Type u_1} {F : Type u_2} [EquivLike F M N] [Mul M] [Mul N] [MulEquivClass F M N] (e : F) {x : M} :
              theorem AddEquivClass.apply_mem_center_iff {M : Type u_3} {N : Type u_1} {F : Type u_2} [EquivLike F M N] [Add M] [Add N] [AddEquivClass F M N] (e : F) {x : M} :
              def Subsemigroup.centerCongr {M : Type u_2} {N : Type u_1} [Mul M] [Mul N] (e : M ≃* N) :
              (center M) ≃* (center N)

              The center of isomorphic magmas are isomorphic.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                def AddSubsemigroup.centerCongr {M : Type u_2} {N : Type u_1} [Add M] [Add N] (e : M ≃+ N) :
                (center M) ≃+ (center N)

                The center of isomorphic additive magmas are isomorphic.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[simp]
                  theorem Subsemigroup.centerCongr_symm_apply_coe {M : Type u_2} {N : Type u_1} [Mul M] [Mul N] (e : M ≃* N) (s : (center N)) :
                  ((centerCongr e).symm s) = e.symm s
                  @[simp]
                  theorem AddSubsemigroup.centerCongr_symm_apply_coe {M : Type u_2} {N : Type u_1} [Add M] [Add N] (e : M ≃+ N) (s : (center N)) :
                  ((centerCongr e).symm s) = e.symm s
                  @[simp]
                  theorem AddSubsemigroup.centerCongr_apply_coe {M : Type u_2} {N : Type u_1} [Add M] [Add N] (e : M ≃+ N) (r : (center M)) :
                  ((centerCongr e) r) = e r
                  @[simp]
                  theorem Subsemigroup.centerCongr_apply_coe {M : Type u_2} {N : Type u_1} [Mul M] [Mul N] (e : M ≃* N) (r : (center M)) :
                  ((centerCongr e) r) = e r
                  def Submonoid.centerCongr {M : Type u_2} {N : Type u_1} [MulOneClass M] [MulOneClass N] (e : M ≃* N) :
                  (center M) ≃* (center N)

                  The center of isomorphic monoids are isomorphic.

                  Equations
                  Instances For
                    def AddSubmonoid.centerCongr {M : Type u_2} {N : Type u_1} [AddZeroClass M] [AddZeroClass N] (e : M ≃+ N) :
                    (center M) ≃+ (center N)

                    The center of isomorphic additive monoids are isomorphic.

                    Equations
                    Instances For
                      @[simp]
                      theorem AddSubmonoid.centerCongr_apply_coe {M : Type u_2} {N : Type u_1} [AddZeroClass M] [AddZeroClass N] (e : M ≃+ N) (r : (AddSubsemigroup.center M)) :
                      ((centerCongr e) r) = e r
                      @[simp]
                      theorem Submonoid.centerCongr_symm_apply_coe {M : Type u_2} {N : Type u_1} [MulOneClass M] [MulOneClass N] (e : M ≃* N) (s : (Subsemigroup.center N)) :
                      ((centerCongr e).symm s) = e.symm s
                      @[simp]
                      theorem AddSubmonoid.centerCongr_symm_apply_coe {M : Type u_2} {N : Type u_1} [AddZeroClass M] [AddZeroClass N] (e : M ≃+ N) (s : (AddSubsemigroup.center N)) :
                      ((centerCongr e).symm s) = e.symm s
                      @[simp]
                      theorem Submonoid.centerCongr_apply_coe {M : Type u_2} {N : Type u_1} [MulOneClass M] [MulOneClass N] (e : M ≃* N) (r : (Subsemigroup.center M)) :
                      ((centerCongr e) r) = e r

                      The center of a magma is isomorphic to the center of its opposite.

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

                        The center of an additive magma is isomorphic to the center of its opposite.

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

                          The center of a monoid is isomorphic to the center of its opposite.

                          Equations
                          Instances For

                            The center of an additive monoid is isomorphic to the center of its opposite.

                            Equations
                            Instances For