Documentation

Mathlib.CategoryTheory.Monoidal.Category

Monoidal categories #

A monoidal category is a category equipped with a tensor product, unitors, and an associator. In the definition, we provide the tensor product as a pair of functions

The tensor product can be expressed as a functor via tensor : C × C ⥤ C. The unitors and associator are gathered together as natural isomorphisms in leftUnitor_nat_iso, rightUnitor_nat_iso and associator_nat_iso.

Some consequences of the definition are proved in other files after proving the coherence theorem, e.g. (λ_ (𝟙_ C)).hom = (ρ_ (𝟙_ C)).hom in CategoryTheory.Monoidal.CoherenceLemmas.

Implementation notes #

In the definition of monoidal categories, we also provide the whiskering operators:

If you want to provide tensorHom and define whiskerLeft and whiskerRight in terms of it, you can use the alternative constructor CategoryTheory.MonoidalCategory.ofTensorHom.

The whiskerings are useful when considering simp-normal forms of morphisms in monoidal categories.

Simp-normal form for morphisms #

Rewriting involving associators and unitors could be very complicated. We try to ease this complexity by putting carefully chosen simp lemmas that rewrite any morphisms into the simp-normal form defined below. Rewriting into simp-normal form is especially useful in preprocessing performed by the coherence tactic.

The simp-normal form of morphisms is defined to be an expression that has the minimal number of parentheses. More precisely,

  1. it is a composition of morphisms like f₁ ≫ f₂ ≫ f₃ ≫ f₄ ≫ f₅ such that each fᵢ is either a structural morphisms (morphisms made up only of identities, associators, unitors) or non-structural morphisms, and
  2. each non-structural morphism in the composition is of the form X₁ ◁ X₂ ◁ X₃ ◁ f ▷ X₄ ▷ X₅, where each Xᵢ is a object that is not the identity or a tensor and f is a non-structural morphisms that is not the identity or a composite.

Note that X₁ ◁ X₂ ◁ X₃ ◁ f ▷ X₄ ▷ X₅ is actually X₁ ◁ (X₂ ◁ (X₃ ◁ ((f ▷ X₄) ▷ X₅))).

Currently, the simp lemmas don't rewrite 𝟙 X ⊗ₘ f and f ⊗ₘ 𝟙 Y into X ◁ f and f ▷ Y, respectively, since it requires a huge refactoring. We hope to add these simp lemmas soon.

References #

Auxiliary structure to carry only the data fields of (and provide notation for) MonoidalCategory.

  • tensorObj : CCC

    curried tensor product of objects

  • whiskerLeft (X : C) {Y₁ Y₂ : C} (f : Y₁ Y₂) : tensorObj X Y₁ tensorObj X Y₂

    left whiskering for morphisms

  • whiskerRight {X₁ X₂ : C} (f : X₁ X₂) (Y : C) : tensorObj X₁ Y tensorObj X₂ Y

    right whiskering for morphisms

  • tensorHom {X₁ Y₁ X₂ Y₂ : C} (f : X₁ Y₁) (g : X₂ Y₂) : tensorObj X₁ X₂ tensorObj Y₁ Y₂

    Tensor product of identity maps is the identity: 𝟙 X₁ ⊗ₘ 𝟙 X₂ = 𝟙 (X₁ ⊗ X₂)

  • tensorUnit : C

    The tensor unity in the monoidal structure 𝟙_ C

  • associator (X Y Z : C) : tensorObj (tensorObj X Y) Z tensorObj X (tensorObj Y Z)

    The associator isomorphism (X ⊗ Y) ⊗ Z ≃ X ⊗ (Y ⊗ Z)

  • leftUnitor (X : C) : tensorObj (tensorUnit C) X X

    The left unitor: 𝟙_ C ⊗ X ≃ X

  • rightUnitor (X : C) : tensorObj X (tensorUnit C) X

    The right unitor: X ⊗ 𝟙_ C ≃ X

Instances

    Notation for tensorObj, the tensor product of objects in a monoidal category

    Equations
      Instances For

        Notation for the whiskerLeft operator of monoidal categories

        Equations
          Instances For

            Notation for the whiskerRight operator of monoidal categories

            Equations
              Instances For

                Notation for tensorHom, the tensor product of morphisms in a monoidal category

                Equations
                  Instances For

                    Notation for tensorUnit, the two-sided identity of

                    Equations
                      Instances For

                        Notation for the monoidal associator: (X ⊗ Y) ⊗ Z ≃ X ⊗ (Y ⊗ Z)

                        Equations
                          Instances For

                            Notation for the leftUnitor: 𝟙_C ⊗ X ≃ X

                            Equations
                              Instances For

                                Notation for the rightUnitor: X ⊗ 𝟙_C ≃ X

                                Equations
                                  Instances For

                                    The property that the pentagon relation is satisfied by four objects in a category equipped with a MonoidalCategoryStruct.

                                    Equations
                                      Instances For

                                        In a monoidal category, we can take the tensor product of objects, X ⊗ Y and of morphisms f ⊗ₘ g. Tensor product does not need to be strictly associative on objects, but there is a specified associator, α_ X Y Z : (X ⊗ Y) ⊗ Z ≅ X ⊗ (Y ⊗ Z). There is a tensor unit 𝟙_ C, with specified left and right unitor isomorphisms λ_ X : 𝟙_ C ⊗ X ≅ X and ρ_ X : X ⊗ 𝟙_ C ≅ X. These associators and unitors satisfy the pentagon and triangle equations.

                                        Instances
                                          theorem CategoryTheory.MonoidalCategory.tensorHom_def_assoc {C : Type u} {𝒞 : Category.{v, u} C} [self : MonoidalCategory C] {X₁ Y₁ X₂ Y₂ : C} (f : X₁ Y₁) (g : X₂ Y₂) {Z : C} (h : tensorObj Y₁ Y₂ Z) :
                                          theorem CategoryTheory.MonoidalCategory.tensor_comp_assoc {C : Type u} {𝒞 : Category.{v, u} C} [self : MonoidalCategory C] {X₁ Y₁ Z₁ X₂ Y₂ Z₂ : C} (f₁ : X₁ Y₁) (f₂ : X₂ Y₂) (g₁ : Y₁ Z₁) (g₂ : Y₂ Z₂) {Z : C} (h : tensorObj Z₁ Z₂ Z) :
                                          theorem CategoryTheory.MonoidalCategory.associator_naturality_assoc {C : Type u} {𝒞 : Category.{v, u} C} [self : MonoidalCategory C] {X₁ X₂ X₃ Y₁ Y₂ Y₃ : C} (f₁ : X₁ Y₁) (f₂ : X₂ Y₂) (f₃ : X₃ Y₃) {Z : C} (h : tensorObj Y₁ (tensorObj Y₂ Y₃) Z) :
                                          @[simp]
                                          theorem CategoryTheory.MonoidalCategory.id_tensorHom {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] (X : C) {Y₁ Y₂ : C} (f : Y₁ Y₂) :
                                          @[simp]
                                          theorem CategoryTheory.MonoidalCategory.tensorHom_id {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X₁ X₂ : C} (f : X₁ X₂) (Y : C) :
                                          theorem CategoryTheory.MonoidalCategory.tensorHom_def' {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X₁ Y₁ X₂ Y₂ : C} (f : X₁ Y₁) (g : X₂ Y₂) :
                                          theorem CategoryTheory.MonoidalCategory.tensorHom_def'_assoc {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X₁ Y₁ X₂ Y₂ : C} (f : X₁ Y₁) (g : X₂ Y₂) {Z : C} (h : tensorObj Y₁ Y₂ Z) :
                                          @[simp]
                                          theorem CategoryTheory.MonoidalCategory.whiskerLeft_hom_inv_assoc {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] (X : C) {Y Z : C} (f : Y Z) {Z✝ : C} (h : tensorObj X Y Z✝) :
                                          @[simp]
                                          @[simp]
                                          theorem CategoryTheory.MonoidalCategory.whiskerLeft_inv_hom_assoc {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] (X : C) {Y Z : C} (f : Y Z) {Z✝ : C} (h : tensorObj X Z Z✝) :
                                          @[simp]
                                          @[simp]
                                          theorem CategoryTheory.MonoidalCategory.whiskerLeft_hom_inv'_assoc {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] (X : C) {Y Z : C} (f : Y Z) [IsIso f] {Z✝ : C} (h : tensorObj X Y Z✝) :
                                          @[simp]
                                          theorem CategoryTheory.MonoidalCategory.hom_inv_whiskerRight'_assoc {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y : C} (f : X Y) [IsIso f] (Z : C) {Z✝ : C} (h : tensorObj X Z Z✝) :
                                          @[simp]
                                          theorem CategoryTheory.MonoidalCategory.whiskerLeft_inv_hom'_assoc {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] (X : C) {Y Z : C} (f : Y Z) [IsIso f] {Z✝ : C} (h : tensorObj X Z Z✝) :
                                          @[simp]
                                          theorem CategoryTheory.MonoidalCategory.inv_hom_whiskerRight'_assoc {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y : C} (f : X Y) [IsIso f] (Z : C) {Z✝ : C} (h : tensorObj Y Z Z✝) :
                                          def CategoryTheory.MonoidalCategory.whiskerLeftIso {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] (X : C) {Y Z : C} (f : Y Z) :

                                          The left whiskering of an isomorphism is an isomorphism.

                                          Equations
                                            Instances For
                                              @[simp]
                                              @[simp]
                                              instance CategoryTheory.MonoidalCategory.whiskerLeft_isIso {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] (X : C) {Y Z : C} (f : Y Z) [IsIso f] :
                                              @[simp]
                                              theorem CategoryTheory.MonoidalCategory.inv_whiskerLeft {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] (X : C) {Y Z : C} (f : Y Z) [IsIso f] :
                                              @[simp]
                                              theorem CategoryTheory.MonoidalCategory.whiskerLeftIso_trans {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] (W : C) {X Y Z : C} (f : X Y) (g : Y Z) :
                                              def CategoryTheory.MonoidalCategory.whiskerRightIso {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y : C} (f : X Y) (Z : C) :

                                              The right whiskering of an isomorphism is an isomorphism.

                                              Equations
                                                Instances For
                                                  @[simp]
                                                  @[simp]
                                                  instance CategoryTheory.MonoidalCategory.whiskerRight_isIso {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y : C} (f : X Y) (Z : C) [IsIso f] :
                                                  @[simp]
                                                  theorem CategoryTheory.MonoidalCategory.inv_whiskerRight {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y : C} (f : X Y) (Z : C) [IsIso f] :
                                                  @[simp]
                                                  def CategoryTheory.MonoidalCategory.tensorIso {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y X' Y' : C} (f : X Y) (g : X' Y') :

                                                  The tensor product of two isomorphisms is an isomorphism.

                                                  Equations
                                                    Instances For
                                                      @[simp]
                                                      theorem CategoryTheory.MonoidalCategory.tensorIso_inv {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y X' Y' : C} (f : X Y) (g : X' Y') :
                                                      @[simp]
                                                      theorem CategoryTheory.MonoidalCategory.tensorIso_hom {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y X' Y' : C} (f : X Y) (g : X' Y') :

                                                      Notation for tensorIso, the tensor product of isomorphisms

                                                      Equations
                                                        Instances For
                                                          theorem CategoryTheory.MonoidalCategory.tensorIso_def {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y X' Y' : C} (f : X Y) (g : X' Y') :
                                                          theorem CategoryTheory.MonoidalCategory.tensorIso_def' {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y X' Y' : C} (f : X Y) (g : X' Y') :
                                                          instance CategoryTheory.MonoidalCategory.tensor_isIso {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {W X Y Z : C} (f : W X) [IsIso f] (g : Y Z) [IsIso g] :
                                                          @[simp]
                                                          theorem CategoryTheory.MonoidalCategory.inv_tensor {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {W X Y Z : C} (f : W X) [IsIso f] (g : Y Z) [IsIso g] :
                                                          inv (tensorHom f g) = tensorHom (inv f) (inv g)
                                                          theorem CategoryTheory.MonoidalCategory.whiskerLeft_dite {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {P : Prop} [Decidable P] (X : C) {Y Z : C} (f : P → (Y Z)) (f' : ¬P → (Y Z)) :
                                                          whiskerLeft X (if h : P then f h else f' h) = if h : P then whiskerLeft X (f h) else whiskerLeft X (f' h)
                                                          theorem CategoryTheory.MonoidalCategory.dite_whiskerRight {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {P : Prop} [Decidable P] {X Y : C} (f : P → (X Y)) (f' : ¬P → (X Y)) (Z : C) :
                                                          whiskerRight (if h : P then f h else f' h) Z = if h : P then whiskerRight (f h) Z else whiskerRight (f' h) Z
                                                          theorem CategoryTheory.MonoidalCategory.tensor_dite {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {P : Prop} [Decidable P] {W X Y Z : C} (f : W X) (g : P → (Y Z)) (g' : ¬P → (Y Z)) :
                                                          tensorHom f (if h : P then g h else g' h) = if h : P then tensorHom f (g h) else tensorHom f (g' h)
                                                          theorem CategoryTheory.MonoidalCategory.dite_tensor {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {P : Prop} [Decidable P] {W X Y Z : C} (f : W X) (g : P → (Y Z)) (g' : ¬P → (Y Z)) :
                                                          tensorHom (if h : P then g h else g' h) f = if h : P then tensorHom (g h) f else tensorHom (g' h) f
                                                          @[simp]
                                                          theorem CategoryTheory.MonoidalCategory.whiskerLeft_eqToHom {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] (X : C) {Y Z : C} (f : Y = Z) :
                                                          @[simp]
                                                          theorem CategoryTheory.MonoidalCategory.eqToHom_whiskerRight {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y : C} (f : X = Y) (Z : C) :

                                                          The lemmas in the next section are true by coherence, but we prove them directly as they are used in proving the coherence theorem.

                                                          @[simp]

                                                          We state it as a simp lemma, which is regarded as an involved version of id_whiskerRight X Y : 𝟙 X ▷ Y = 𝟙 (X ⊗ Y).

                                                          @[deprecated CategoryTheory.MonoidalCategory.leftUnitor_tensor_hom (since := "2025-06-24")]

                                                          Alias of CategoryTheory.MonoidalCategory.leftUnitor_tensor_hom.

                                                          @[deprecated CategoryTheory.MonoidalCategory.rightUnitor_tensor_hom (since := "2025-06-24")]

                                                          Alias of CategoryTheory.MonoidalCategory.rightUnitor_tensor_hom.

                                                          theorem CategoryTheory.MonoidalCategory.associator_inv_naturality {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y Z X' Y' Z' : C} (f : X X') (g : Y Y') (h : Z Z') :
                                                          theorem CategoryTheory.MonoidalCategory.associator_inv_naturality_assoc {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y Z X' Y' Z' : C} (f : X X') (g : Y Y') (h : Z Z') {Z✝ : C} (h✝ : tensorObj (tensorObj X' Y') Z' Z✝) :
                                                          @[simp]
                                                          theorem CategoryTheory.MonoidalCategory.associator_conjugation {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X X' Y Y' Z Z' : C} (f : X X') (g : Y Y') (h : Z Z') :
                                                          theorem CategoryTheory.MonoidalCategory.associator_conjugation_assoc {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X X' Y Y' Z Z' : C} (f : X X') (g : Y Y') (h : Z Z') {Z✝ : C} (h✝ : tensorObj (tensorObj X' Y') Z' Z✝) :
                                                          theorem CategoryTheory.MonoidalCategory.associator_inv_conjugation {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X X' Y Y' Z Z' : C} (f : X X') (g : Y Y') (h : Z Z') :
                                                          theorem CategoryTheory.MonoidalCategory.associator_inv_conjugation_assoc {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X X' Y Y' Z Z' : C} (f : X X') (g : Y Y') (h : Z Z') {Z✝ : C} (h✝ : tensorObj X' (tensorObj Y' Z') Z✝) :
                                                          @[simp]
                                                          @[simp]
                                                          @[simp]
                                                          @[simp]
                                                          @[reducible, inline]
                                                          abbrev CategoryTheory.MonoidalCategory.ofTensorHom {C : Type u} [𝒞 : Category.{v, u} C] [MonoidalCategoryStruct C] (id_tensorHom_id : ∀ (X₁ X₂ : C), tensorHom (CategoryStruct.id X₁) (CategoryStruct.id X₂) = CategoryStruct.id (tensorObj X₁ X₂) := by cat_disch) (id_tensorHom : ∀ (X : C) {Y₁ Y₂ : C} (f : Y₁ Y₂), tensorHom (CategoryStruct.id X) f = whiskerLeft X f := by cat_disch) (tensorHom_id : ∀ {X₁ X₂ : C} (f : X₁ X₂) (Y : C), tensorHom f (CategoryStruct.id Y) = whiskerRight f Y := by cat_disch) (tensor_comp : ∀ {X₁ Y₁ Z₁ X₂ Y₂ Z₂ : C} (f₁ : X₁ Y₁) (f₂ : X₂ Y₂) (g₁ : Y₁ Z₁) (g₂ : Y₂ Z₂), tensorHom (CategoryStruct.comp f₁ g₁) (CategoryStruct.comp f₂ g₂) = CategoryStruct.comp (tensorHom f₁ f₂) (tensorHom g₁ g₂) := by cat_disch) (associator_naturality : ∀ {X₁ X₂ X₃ Y₁ Y₂ Y₃ : C} (f₁ : X₁ Y₁) (f₂ : X₂ Y₂) (f₃ : X₃ Y₃), CategoryStruct.comp (tensorHom (tensorHom f₁ f₂) f₃) (associator Y₁ Y₂ Y₃).hom = CategoryStruct.comp (associator X₁ X₂ X₃).hom (tensorHom f₁ (tensorHom f₂ f₃)) := by cat_disch) (leftUnitor_naturality : ∀ {X Y : C} (f : X Y), CategoryStruct.comp (tensorHom (CategoryStruct.id (tensorUnit C)) f) (leftUnitor Y).hom = CategoryStruct.comp (leftUnitor X).hom f := by cat_disch) (rightUnitor_naturality : ∀ {X Y : C} (f : X Y), CategoryStruct.comp (tensorHom f (CategoryStruct.id (tensorUnit C))) (rightUnitor Y).hom = CategoryStruct.comp (rightUnitor X).hom f := by cat_disch) (pentagon : ∀ (W X Y Z : C), CategoryStruct.comp (tensorHom (associator W X Y).hom (CategoryStruct.id Z)) (CategoryStruct.comp (associator W (tensorObj X Y) Z).hom (tensorHom (CategoryStruct.id W) (associator X Y Z).hom)) = CategoryStruct.comp (associator (tensorObj W X) Y Z).hom (associator W X (tensorObj Y Z)).hom := by cat_disch) (triangle : ∀ (X Y : C), CategoryStruct.comp (associator X (tensorUnit C) Y).hom (tensorHom (CategoryStruct.id X) (leftUnitor Y).hom) = tensorHom (rightUnitor X).hom (CategoryStruct.id Y) := by cat_disch) :

                                                          A constructor for monoidal categories that requires tensorHom instead of whiskerLeft and whiskerRight.

                                                          Equations
                                                            Instances For

                                                              The tensor product expressed as a functor.

                                                              Equations
                                                                Instances For
                                                                  @[simp]
                                                                  @[simp]
                                                                  theorem CategoryTheory.MonoidalCategory.tensor_map (C : Type u) [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y : C × C} (f : X Y) :
                                                                  (tensor C).map f = tensorHom f.1 f.2

                                                                  The left-associated triple tensor product as a functor.

                                                                  Equations
                                                                    Instances For
                                                                      @[simp]
                                                                      theorem CategoryTheory.MonoidalCategory.leftAssocTensor_map (C : Type u) [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y : C × C × C} (f : X Y) :
                                                                      (leftAssocTensor C).map f = tensorHom (tensorHom f.1 f.2.1) f.2.2

                                                                      The right-associated triple tensor product as a functor.

                                                                      Equations
                                                                        Instances For
                                                                          @[simp]
                                                                          theorem CategoryTheory.MonoidalCategory.rightAssocTensor_map (C : Type u) [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X Y : C × C × C} (f : X Y) :
                                                                          (rightAssocTensor C).map f = tensorHom f.1 (tensorHom f.2.1 f.2.2)

                                                                          The tensor product bifunctor C ⥤ C ⥤ C of a monoidal category.

                                                                          Equations
                                                                            Instances For
                                                                              @[simp]
                                                                              theorem CategoryTheory.MonoidalCategory.curriedTensor_obj_map (C : Type u) [𝒞 : Category.{v, u} C] [MonoidalCategory C] (X : C) {X✝ Y✝ : C} (g : X✝ Y✝) :
                                                                              @[simp]
                                                                              theorem CategoryTheory.MonoidalCategory.curriedTensor_map_app (C : Type u) [𝒞 : Category.{v, u} C] [MonoidalCategory C] {X✝ Y✝ : C} (f : X✝ Y✝) (Y : C) :
                                                                              @[reducible, inline]

                                                                              Tensoring on the left with a fixed object, as a functor.

                                                                              Equations
                                                                                Instances For
                                                                                  @[reducible, inline]

                                                                                  Tensoring on the right with a fixed object, as a functor.

                                                                                  Equations
                                                                                    Instances For
                                                                                      @[reducible, inline]

                                                                                      The functor fun X ↦ 𝟙_ C ⊗ X.

                                                                                      Equations
                                                                                        Instances For
                                                                                          @[reducible, inline]

                                                                                          The functor fun X ↦ X ⊗ 𝟙_ C.

                                                                                          Equations
                                                                                            Instances For

                                                                                              The associator as a natural isomorphism.

                                                                                              Equations
                                                                                                Instances For

                                                                                                  The left unitor as a natural isomorphism.

                                                                                                  Equations
                                                                                                    Instances For

                                                                                                      The right unitor as a natural isomorphism.

                                                                                                      Equations
                                                                                                        Instances For

                                                                                                          The associator as a natural isomorphism between trifunctors C ⥤ C ⥤ C ⥤ C.

                                                                                                          Equations
                                                                                                            Instances For

                                                                                                              Tensoring on the left with X ⊗ Y is naturally isomorphic to tensoring on the left with Y, and then again with X.

                                                                                                              Equations
                                                                                                                Instances For
                                                                                                                  @[reducible, inline]

                                                                                                                  Tensoring on the left, as a functor from C into endofunctors of C.

                                                                                                                  TODO: show this is an op-monoidal functor.

                                                                                                                  Equations
                                                                                                                    Instances For
                                                                                                                      @[reducible, inline]

                                                                                                                      Tensoring on the right, as a functor from C into endofunctors of C.

                                                                                                                      We later show this is a monoidal functor.

                                                                                                                      Equations
                                                                                                                        Instances For

                                                                                                                          Tensoring on the right with X ⊗ Y is naturally isomorphic to tensoring on the right with X, and then again with Y.

                                                                                                                          Equations
                                                                                                                            Instances For
                                                                                                                              @[simp]
                                                                                                                              theorem CategoryTheory.MonoidalCategory.prodMonoidal_whiskerRight (C₁ : Type u₁) [Category.{v₁, u₁} C₁] [MonoidalCategory C₁] (C₂ : Type u₂) [Category.{v₂, u₂} C₂] [MonoidalCategory C₂] {X₁✝ X₂✝ : C₁ × C₂} (f : X₁✝ X₂✝) (X : C₁ × C₂) :
                                                                                                                              @[simp]
                                                                                                                              theorem CategoryTheory.MonoidalCategory.prodMonoidal_tensorHom (C₁ : Type u₁) [Category.{v₁, u₁} C₁] [MonoidalCategory C₁] (C₂ : Type u₂) [Category.{v₂, u₂} C₂] [MonoidalCategory C₂] {X₁✝ Y₁✝ X₂✝ Y₂✝ : C₁ × C₂} (f : X₁✝ Y₁✝) (g : X₂✝ Y₂✝) :
                                                                                                                              tensorHom f g = (tensorHom f.1 g.1, tensorHom f.2 g.2)
                                                                                                                              @[simp]
                                                                                                                              theorem CategoryTheory.MonoidalCategory.prodMonoidal_associator (C₁ : Type u₁) [Category.{v₁, u₁} C₁] [MonoidalCategory C₁] (C₂ : Type u₂) [Category.{v₂, u₂} C₂] [MonoidalCategory C₂] (X Y Z : C₁ × C₂) :
                                                                                                                              associator X Y Z = (associator X.1 Y.1 Z.1).prod (associator X.2 Y.2 Z.2)
                                                                                                                              @[simp]
                                                                                                                              theorem CategoryTheory.MonoidalCategory.prodMonoidal_tensorObj (C₁ : Type u₁) [Category.{v₁, u₁} C₁] [MonoidalCategory C₁] (C₂ : Type u₂) [Category.{v₂, u₂} C₂] [MonoidalCategory C₂] (X Y : C₁ × C₂) :
                                                                                                                              tensorObj X Y = (tensorObj X.1 Y.1, tensorObj X.2 Y.2)
                                                                                                                              @[simp]
                                                                                                                              theorem CategoryTheory.MonoidalCategory.prodMonoidal_whiskerLeft (C₁ : Type u₁) [Category.{v₁, u₁} C₁] [MonoidalCategory C₁] (C₂ : Type u₂) [Category.{v₂, u₂} C₂] [MonoidalCategory C₂] (X x✝ x✝¹ : C₁ × C₂) (f : x✝ x✝¹) :
                                                                                                                              @[reducible, inline]
                                                                                                                              noncomputable abbrev CategoryTheory.MonoidalCategory.fullSubcategory {C : Type u} [Category.{v, u} C] [MonoidalCategory C] (P : ObjectProperty C) (tensorUnit : P (tensorUnit C)) (tensorObj : ∀ (X Y : C), P XP YP (tensorObj X Y)) :

                                                                                                                              The restriction of a monoidal category along an object property that's closed under the monoidal structure.

                                                                                                                              Equations
                                                                                                                                Instances For