Documentation

Mathlib.CategoryTheory.Monoidal.Multifunctor

Constructing monoidal functors from natural transformations between multifunctors #

This file provides alternative constructors for (op/lax) monoidal functors, given tensorators μ : F - ⊗ F - ⟶ F (- ⊗ -) / δ : F (- ⊗ -) ⟶ F - ⊗ F - as natural transformations between bifunctors. The associativity conditions are phrased as equalities of natural transformations between trifunctors (F - ⊗ F -) ⊗ F - ⟶ F (- ⊗ (- ⊗ -)) / F ((- ⊗ -) ⊗ -) ⟶ F - ⊗ (F - ⊗ F -), and the unitality conditions are phrased as equalities of natural transformation between functors.

Once we have more API for quadrifunctors, we can add constructors for monoidal category structures by phrasing the pentagon axiom as an equality of natural transformations between quadrifunctors.

@[reducible, inline]

The bifunctor (F -) ⊗ -.

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

    The bifunctor - ⊗ (F -).

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

      The trifunctor (F - ⊗ F -) ⊗ F -.

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

        The trifunctor F - ⊗ (F - ⊗ F -).

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

          Lax monoidal functors #

          Given a unit morphism ε : 𝟙_ D ⟶ F.obj (𝟙_ C)) and a tensorator μ : F - ⊗ F - ⟶ F (- ⊗ -) such that the diagrams below commute, we define CategoryTheory.Functor.LaxMonoidal.ofBifunctor : F.LaxMonoidal.

          Associativity hexagon #

                (F - ⊗ F -) ⊗ F -
                  /           \
                 v             v
          F (- ⊗ -) ⊗ F -    F - ⊗ (F - ⊗ F -)
                 |             |
                 v             v
          F ((- ⊗ -) ⊗ -)    F - ⊗ F (- ⊗ -)
                  \            /
                   v          v
                 F (- ⊗ (- ⊗ -))
          

          Left unitality square #

          𝟙 ⊗ F - ⟶ F 𝟙 ⊗ F -
            |           |
            v           v
            F    ←   F (𝟙 ⊗ -)
          

          Right unitality square #

          F - ⊗ 𝟙 ⟶ F - ⊗ F 𝟙
            |           |
            v           v
            F   ←   F (- ⊗ 𝟙)
          

          The composition of the left maps in the associativity hexagon.

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

            The top right map in the associativity hexagon.

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

              The composition of the right maps in the associativity hexagon.

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

                F is lax monoidal given a unit morphism ε : 𝟙_ D ⟶ F.obj (𝟙_ C)) and a tensorator μ : F - ⊗ F - ⟶ F (- ⊗ -) as a natural transformation between bifunctors, satisfying the relevant compatibilities.

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

                  Oplax monoidal functors #

                  Given a counit morphism η : F.obj (𝟙_ C)) ⟶ 𝟙_ D and a tensorator δ : F (- ⊗ -) ⟶ F - ⊗ F - such that the diagrams below commute, we define CategoryTheory.Functor.OplaxMonoidal.ofBifunctor : F.OplaxMonoidal.

                  Oplax associativity hexagon #

                        F ((- ⊗ -) ⊗ -)
                          /           \
                         v             v
                  F (- ⊗ -) ⊗ F -      F (- ⊗ (- ⊗ -))
                         |                |
                         v                v
                  (F - ⊗ F -) ⊗ F -    F - ⊗ F (- ⊗ -)
                          \            /
                           v          v
                         F - ⊗ (F - ⊗ F -)
                  

                  Oplax left unitality square #

                    F   ⟶  F (𝟙 ⊗ -)
                    |           |
                    v           v
                  𝟙 ⊗ F - ← F 𝟙 ⊗ F -
                  

                  Oplax right unitality square #

                    F  ⟶   F (- ⊗ 𝟙)
                    |           |
                    v           v
                  F - ⊗ 𝟙 ← F - ⊗ F 𝟙
                  

                  The bottom left map in the oplax associativity hexagon.

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

                    The composition of the three left maps in the oplax associativity hexagon.

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

                      The composition of the three right maps in the oplax associativity hexagon.

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

                        F is oplax monoidal given a counit morphism η : F.obj (𝟙_ C) ⟶ 𝟙_ D and a tensorator δ : F (- ⊗ -) ⟶ F - ⊗ F - as a natural transformation between bifunctors, satisfying the relevant compatibilities.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          def CategoryTheory.Functor.Monoidal.ofBifunctor {C : Type u_1} [Category.{u_3, u_1} C] [MonoidalCategory C] {D : Type u_2} [Category.{u_4, u_2} D] [MonoidalCategory D] {F : Functor C D} (ε : MonoidalCategoryStruct.tensorUnit D F.obj (MonoidalCategoryStruct.tensorUnit C)) (μ : MonoidalCategory.curriedTensorPre F MonoidalCategory.curriedTensorPost F) (associativity : LaxMonoidal.ofBifunctor.firstMap μ = LaxMonoidal.ofBifunctor.secondMap μ) (left_unitality : LaxMonoidal.ofBifunctor.leftMapₗ F = CategoryStruct.comp (LaxMonoidal.ofBifunctor.topMapₗ ε) (CategoryStruct.comp (μ.app (MonoidalCategoryStruct.tensorUnit C)) (LaxMonoidal.ofBifunctor.bottomMapₗ F))) (right_unitality : LaxMonoidal.ofBifunctor.leftMapᵣ F = CategoryStruct.comp (LaxMonoidal.ofBifunctor.topMapᵣ ε) (CategoryStruct.comp (((flipFunctor C C D).map μ).app (MonoidalCategoryStruct.tensorUnit C)) (LaxMonoidal.ofBifunctor.bottomMapᵣ F))) (η : F.obj (MonoidalCategoryStruct.tensorUnit C) MonoidalCategoryStruct.tensorUnit D) (δ : MonoidalCategory.curriedTensorPost F MonoidalCategory.curriedTensorPre F) (oplax_associativity : OplaxMonoidal.ofBifunctor.firstMap δ = OplaxMonoidal.ofBifunctor.secondMap δ) (oplax_left_unitality : OplaxMonoidal.ofBifunctor.leftMapₗ F = CategoryStruct.comp (OplaxMonoidal.ofBifunctor.topMapₗ F) (CategoryStruct.comp (δ.app (MonoidalCategoryStruct.tensorUnit C)) (OplaxMonoidal.ofBifunctor.bottomMapₗ η))) (oplax_right_unitality : OplaxMonoidal.ofBifunctor.leftMapᵣ F = CategoryStruct.comp (OplaxMonoidal.ofBifunctor.topMapᵣ F) (CategoryStruct.comp (((flipFunctor C C D).map δ).app (MonoidalCategoryStruct.tensorUnit C)) (OplaxMonoidal.ofBifunctor.bottomMapᵣ η))) (ε_η : CategoryStruct.comp ε η = CategoryStruct.id (MonoidalCategoryStruct.tensorUnit D)) (η_ε : CategoryStruct.comp η ε = CategoryStruct.id (F.obj (MonoidalCategoryStruct.tensorUnit C))) (μ_δ : CategoryStruct.comp μ δ = CategoryStruct.id (MonoidalCategory.curriedTensorPre F)) (δ_μ : CategoryStruct.comp δ μ = CategoryStruct.id (MonoidalCategory.curriedTensorPost F)) :

                          F is monoidal given a co/unit morphisms ε/η : 𝟙_ D ↔ F.obj (𝟙_ C) and tensorators μ / δ : F - ⊗ F - ↔ F (- ⊗ -) as natural transformations between bifunctors, satisfying the relevant compatibilities.

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

                            F is monoidal given a unit isomorphism ε : 𝟙_ D ≅ F.obj (𝟙_ C) and a tensorator isomorphism μ : F - ⊗ F - ≅ F (- ⊗ -) as a natural isomorphism between bifunctors, satisfying the relevant compatibilities.

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