Documentation

Mathlib.CategoryTheory.Monoidal.Mon_

The category of monoids in a monoidal category. #

We define monoids in a monoidal category C and show that the category of monoids is equivalent to the category of lax monoidal functors from the unit monoidal category to C. We also show that if C is braided, then the category of monoids is naturally monoidal.

Simp set for monoid object tautologies #

In this file, we also provide a simp set called mon_tauto whose goal is to prove all tautologies about morphisms from some (tensor) power of M to M, where M is a (commutative) monoid object in a (braided) monoidal category.

Please read the documentation in Mathlib/Tactic/Attr/Register.lean for full details.

TODO #

A monoid object internal to a monoidal category.

When the monoidal category is preadditive, this is also sometimes called an "algebra object".

Instances
    @[deprecated CategoryTheory.MonObj (since := "2025-09-09")]

    Alias of CategoryTheory.MonObj.


    A monoid object internal to a monoidal category.

    When the monoidal category is preadditive, this is also sometimes called an "algebra object".

    Equations
    Instances For

      The multiplication morphism of a monoid object.

      Equations
      Instances For

        The multiplication morphism of a monoid object.

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

          The unit morphism of a monoid object.

          Equations
          Instances For

            The unit morphism of a monoid object.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              def CategoryTheory.MonObj.ofIso {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M X : C} [MonObj M] (e : M X) :

              Transfer MonObj along an isomorphism.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                Equations
                • One or more equations did not get rendered due to their size.
                theorem CategoryTheory.MonObj.ext {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {X : C} (h₁ h₂ : MonObj X) (H : mul = mul) :
                h₁ = h₂
                theorem CategoryTheory.MonObj.ext_iff {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {X : C} {h₁ h₂ : MonObj X} :
                h₁ = h₂ mul = mul
                class CategoryTheory.IsMonHom {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M N : C} [MonObj M] [MonObj N] (f : M N) :

                The property that a morphism between monoid objects is a monoid morphism.

                Instances
                  @[deprecated CategoryTheory.IsMonHom (since := "2025-09-15")]
                  def CategoryTheory.IsMon_Hom {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M N : C} [MonObj M] [MonObj N] (f : M N) :

                  Alias of CategoryTheory.IsMonHom.


                  The property that a morphism between monoid objects is a monoid morphism.

                  Equations
                  Instances For
                    @[simp]
                    theorem CategoryTheory.IsMonHom.mul_hom_assoc {C : Type u₁} {inst✝ : Category.{v₁, u₁} C} {inst✝¹ : MonoidalCategory C} {M N : C} {inst✝² : MonObj M} {inst✝³ : MonObj N} (f : M N) [self : IsMonHom f] {Z : C} (h : N Z) :
                    @[simp]
                    theorem CategoryTheory.IsMonHom.one_hom_assoc {C : Type u₁} {inst✝ : Category.{v₁, u₁} C} {inst✝¹ : MonoidalCategory C} {M N : C} {inst✝² : MonObj M} {inst✝³ : MonObj N} (f : M N) [self : IsMonHom f] {Z : C} (h : N Z) :
                    instance CategoryTheory.instIsMonHomComp {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M N O : C} [MonObj M] [MonObj N] [MonObj O] (f : M N) (g : N O) [IsMonHom f] [IsMonHom g] :
                    structure CategoryTheory.Mon (C : Type u₁) [Category.{v₁, u₁} C] [MonoidalCategory C] :
                    Type (max u₁ v₁)

                    A monoid object internal to a monoidal category.

                    When the monoidal category is preadditive, this is also sometimes called an "algebra object".

                    • X : C

                      The underlying object in the ambient monoidal category

                    • mon : MonObj self.X
                    Instances For
                      @[deprecated CategoryTheory.Mon (since := "2025-09-15")]
                      def CategoryTheory.Mon_ (C : Type u₁) [Category.{v₁, u₁} C] [MonoidalCategory C] :
                      Type (max u₁ v₁)

                      Alias of CategoryTheory.Mon.


                      A monoid object internal to a monoidal category.

                      When the monoidal category is preadditive, this is also sometimes called an "algebra object".

                      Equations
                      Instances For

                        The trivial monoid object. We later show this is initial in Mon C.

                        Equations
                        Instances For

                          In this section, we prove that the category of monoids in a braided monoidal category is monoidal.

                          Given two monoids M and N in a braided monoidal category C, the multiplication on the tensor product M.X ⊗ N.X is defined in the obvious way: it is the tensor product of the multiplications on M and N, except that the tensor factors in the source come in the wrong order, which we fix by pre-composing with a permutation isomorphism constructed from the braiding.

                          (There is a subtlety here: in fact there are two ways to do these, using either the positive or negative crossing.)

                          A more conceptual way of understanding this definition is the following: The braiding on C gives rise to a monoidal structure on the tensor product functor from C × C to C. A pair of monoids in C gives rise to a monoid in C × C, which the tensor product functor by being monoidal takes to a monoid in C. The permutation isomorphism appearing in the definition of the multiplication on the tensor product of two monoids is an instance of a more general family of isomorphisms which together form a strength that equips the tensor product functor with a monoidal structure, and the monoid axioms for the tensor product follow from the monoid axioms for the tensor factors plus the properties of the strength (i.e., monoidal functor axioms). The strength tensorμ of the tensor product functor has been defined in Mathlib/CategoryTheory/Monoidal/Braided.lean. Its properties, stated as independent lemmas in that module, are used extensively in the proofs below. Notice that we could have followed the above plan not only conceptually but also as a possible implementation and could have constructed the tensor product of monoids via mapMon, but we chose to give a more explicit definition directly in terms of tensorμ.

                          To complete the definition of the monoidal category structure on the category of monoids, we need to provide definitions of associator and unitors. The obvious candidates are the associator and unitors from C, but we need to prove that they are monoid morphisms, i.e., compatible with unit and multiplication. These properties translate to the monoidality of the associator and unitors (with respect to the monoidal structures on the functors they relate), which have also been proved in Mathlib/CategoryTheory/Monoidal/Braided.lean.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          structure CategoryTheory.Mon.Hom {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] (M N : Mon C) :
                          Type v₁

                          A morphism of monoid objects.

                          Instances For
                            theorem CategoryTheory.Mon.Hom.ext_iff {C : Type u₁} {inst✝ : Category.{v₁, u₁} C} {inst✝¹ : MonoidalCategory C} {M N : Mon C} {x y : M.Hom N} :
                            x = y x.hom = y.hom
                            theorem CategoryTheory.Mon.Hom.ext {C : Type u₁} {inst✝ : Category.{v₁, u₁} C} {inst✝¹ : MonoidalCategory C} {M N : Mon C} {x y : M.Hom N} (hom : x.hom = y.hom) :
                            x = y
                            @[reducible, inline]

                            Construct a morphism M ⟶ N of Mon C from a map f : M ⟶ N and a IsMonHom f instance.

                            Equations
                            Instances For

                              The identity morphism on a monoid object.

                              Equations
                              Instances For
                                Equations
                                def CategoryTheory.Mon.comp {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M N O : Mon C} (f : M.Hom N) (g : N.Hom O) :
                                M.Hom O

                                Composition of morphisms of monoid objects.

                                Equations
                                Instances For
                                  @[simp]
                                  theorem CategoryTheory.Mon.comp_hom {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M N O : Mon C} (f : M.Hom N) (g : N.Hom O) :
                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  theorem CategoryTheory.Mon.Hom.ext' {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M N : Mon C} {f g : M N} (w : f.hom = g.hom) :
                                  f = g
                                  theorem CategoryTheory.Mon.Hom.ext'_iff {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M N : Mon C} {f g : M N} :
                                  f = g f.hom = g.hom
                                  @[simp]

                                  The forgetful functor from monoid objects to the ambient category.

                                  Equations
                                  Instances For
                                    @[simp]
                                    theorem CategoryTheory.Mon.forget_map (C : Type u₁) [Category.{v₁, u₁} C] [MonoidalCategory C] {X✝ Y✝ : Mon C} (f : X✝ Y✝) :
                                    (forget C).map f = f.hom
                                    @[simp]

                                    The forgetful functor from monoid objects to the ambient category reflects isomorphisms.

                                    def CategoryTheory.Mon.mkIso' {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M N : C} [MonObj M] [MonObj N] (e : M N) [IsMonHom e.hom] :
                                    { X := M, mon := inst✝ } { X := N, mon := inst✝¹ }

                                    Construct an isomorphism of monoid objects by giving a monoid isomorphism between the underlying objects.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem CategoryTheory.Mon.mkIso'_hom_hom {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M N : C} [MonObj M] [MonObj N] (e : M N) [IsMonHom e.hom] :
                                      @[simp]
                                      theorem CategoryTheory.Mon.mkIso'_inv_hom {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {M N : C} [MonObj M] [MonObj N] (e : M N) [IsMonHom e.hom] :
                                      @[reducible, inline]

                                      Construct an isomorphism of monoid objects by giving an isomorphism between the underlying objects and checking compatibility with unit and multiplication only in the forward direction.

                                      Equations
                                      Instances For
                                        Equations
                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        @[simp]
                                        theorem CategoryTheory.Mon.monMonoidalStruct_tensorHom_hom {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] {X₁✝ Y₁✝ X₂✝ Y₂✝ : Mon C} (f : X₁✝ Y₁✝) (g : X₂✝ Y₂✝) :
                                        Equations
                                        • One or more equations did not get rendered due to their size.

                                        The forgetful functor from Mon C to C is monoidal when C is monoidal.

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

                                        We next show that if C is symmetric, then Mon C is braided, and indeed symmetric.

                                        Note that Mon C is not braided in general when C is only braided.

                                        The more interesting construction is the 2-category of monoids in C, bimodules between the monoids, and intertwiners between the bimodules.

                                        When C is braided, that is a monoidal 2-category.

                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        @[reducible, inline]

                                        The image of a monoid object under a lax monoidal functor is a monoid object.

                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        Instances For
                                          @[deprecated CategoryTheory.Functor.monObjObj (since := "2025-09-09")]

                                          Alias of CategoryTheory.Functor.monObjObj.


                                          The image of a monoid object under a lax monoidal functor is a monoid object.

                                          Equations
                                          Instances For

                                            A lax monoidal functor takes monoid objects to monoid objects.

                                            That is, a lax monoidal functor F : C ⥤ D induces a functor Mon C ⥤ Mon D.

                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For
                                              @[simp]
                                              theorem CategoryTheory.Functor.mapMon_map_hom {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] {D : Type u₂} [Category.{v₂, u₂} D] [MonoidalCategory D] (F : Functor C D) [F.LaxMonoidal] {X✝ Y✝ : Mon C} (f : X✝ Y✝) :
                                              (F.mapMon.map f).hom = F.map f.hom

                                              The identity functor is also the identity on monoid objects.

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

                                                The composition functor is also the composition on monoid objects.

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

                                                  Natural transformations between functors lift to monoid objects.

                                                  Equations
                                                  Instances For

                                                    Natural isomorphisms between functors lift to monoid objects.

                                                    Equations
                                                    Instances For
                                                      @[reducible, inline]

                                                      Pullback a monoid object along a fully faithful oplax monoidal functor.

                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For
                                                        @[deprecated CategoryTheory.Functor.FullyFaithful.monObj (since := "2025-09-09")]

                                                        Alias of CategoryTheory.Functor.FullyFaithful.monObj.


                                                        Pullback a monoid object along a fully faithful oplax monoidal functor.

                                                        Equations
                                                        Instances For

                                                          If F : C ⥤ D is a fully faithful monoidal functor, then Mon(F) : Mon C ⥤ Mon D is fully faithful too.

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

                                                            The essential image of a fully faithful functor between cartesian-monoidal categories is the same on monoid objects as on objects.

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

                                                            mapMon is functorial in the lax monoidal functor.

                                                            Equations
                                                            • One or more equations did not get rendered due to their size.
                                                            Instances For
                                                              @[simp]
                                                              theorem CategoryTheory.Functor.mapMonFunctor_map_app_hom (C : Type u₁) [Category.{v₁, u₁} C] [MonoidalCategory C] (D : Type u₂) [Category.{v₂, u₂} D] [MonoidalCategory D] {X✝ Y✝ : LaxMonoidalFunctor C D} (α : X✝ Y✝) (A : Mon C) :
                                                              (((mapMonFunctor C D).map α).app A).hom = α.hom.app A.X

                                                              An adjunction of monoidal functors lifts to an adjunction of their lifts to monoid objects.

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

                                                                An equivalence of categories lifts to an equivalence of their monoid objects.

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

                                                                  Implementation of Mon.equivLaxMonoidalFunctorPUnit.

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

                                                                    Implementation of Mon.equivLaxMonoidalFunctorPUnit.

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

                                                                      Implementation of Mon.equivLaxMonoidalFunctorPUnit.

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

                                                                        Auxiliary definition for counitIso.

                                                                        Equations
                                                                        • One or more equations did not get rendered due to their size.
                                                                        Instances For
                                                                          @[deprecated CategoryTheory.Mon.EquivLaxMonoidalFunctorPUnit.isMonHom_counitIsoAux (since := "2025-09-15")]

                                                                          Alias of CategoryTheory.Mon.EquivLaxMonoidalFunctorPUnit.isMonHom_counitIsoAux.

                                                                          Implementation of Mon.equivLaxMonoidalFunctorPUnit.

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

                                                                            Monoid objects in C are "just" lax monoidal functors from the trivial monoidal category to C.

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

                                                                              Predicate for a monoid object to be commutative.

                                                                              Instances
                                                                                @[deprecated CategoryTheory.IsCommMonObj (since := "2025-09-14")]

                                                                                Alias of CategoryTheory.IsCommMonObj.


                                                                                Predicate for a monoid object to be commutative.

                                                                                Equations
                                                                                Instances For
                                                                                  @[simp]
                                                                                  theorem CategoryTheory.IsCommMonObj.mul_comm_assoc {C : Type u₁} {inst✝ : Category.{v₁, u₁} C} {inst✝¹ : MonoidalCategory C} {inst✝² : BraidedCategory C} (X : C) {inst✝³ : MonObj X} [self : IsCommMonObj X] {Z : C} (h : X Z) :