Documentation

Mathlib.CategoryTheory.Monoidal.Braided.Basic

Braided and symmetric monoidal categories #

The basic definitions of braided monoidal categories, and symmetric monoidal categories, as well as braided functors.

Implementation note #

We make BraidedCategory another typeclass, but then have SymmetricCategory extend this. The rationale is that we are not carrying any additional data, just requiring a property.

Future work #

References #

A braided monoidal category is a monoidal category equipped with a braiding isomorphism β_ X Y : X ⊗ Y ≅ Y ⊗ X which is natural in both arguments, and also satisfies the two hexagon identities.

Instances

    The braiding natural isomorphism.

    Equations
      Instances For

        In a braided monoidal category, the functors tensorLeft X and tensorRight X are isomorphic.

        Equations
          Instances For

            Verifying the axioms for a braiding by checking that the candidate braiding is sent to a braiding by a faithful monoidal functor.

            Equations
              Instances For
                @[deprecated CategoryTheory.BraidedCategory.ofFaithful (since := "2025-07-12")]

                Alias of CategoryTheory.BraidedCategory.ofFaithful.


                Verifying the axioms for a braiding by checking that the candidate braiding is sent to a braiding by a faithful monoidal functor.

                Equations
                  Instances For

                    Pull back a braiding along a fully faithful monoidal functor.

                    Equations
                      Instances For
                        @[deprecated CategoryTheory.BraidedCategory.ofFullyFaithful (since := "2025-07-12")]

                        Alias of CategoryTheory.BraidedCategory.ofFullyFaithful.


                        Pull back a braiding along a fully faithful monoidal functor.

                        Equations
                          Instances For

                            We now establish how the braiding interacts with the unitors.

                            I couldn't find a detailed proof in print, but this is discussed in:

                            A symmetric monoidal category is a braided monoidal category for which the braiding is symmetric.

                            Instances

                              A lax braided functor between braided monoidal categories is a lax monoidal functor which preserves the braiding.

                              Instances
                                theorem CategoryTheory.Functor.LaxBraided.braided_assoc {C : Type u₁} {inst✝ : Category.{v₁, u₁} C} {inst✝¹ : MonoidalCategory C} {inst✝² : BraidedCategory C} {D : Type u₂} {inst✝³ : Category.{v₂, u₂} D} {inst✝⁴ : MonoidalCategory D} {inst✝⁵ : BraidedCategory D} {F : Functor C D} [self : F.LaxBraided] (X Y : C) {Z : D} (h : F.obj (MonoidalCategoryStruct.tensorObj Y X) Z) :
                                structure CategoryTheory.LaxBraidedFunctor (C : Type u₁) [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (D : Type u₂) [Category.{v₂, u₂} D] [MonoidalCategory D] [BraidedCategory D] extends CategoryTheory.Functor C D :
                                Type (max (max (max u₁ u₂) v₁) v₂)

                                Bundled version of lax braided functors.

                                Instances For

                                  The lax monoidal functor induced by a lax braided functor.

                                  Equations
                                    Instances For

                                      Constructor for morphisms in the category LaxBraiededFunctor C D.

                                      Equations
                                        Instances For

                                          Constructor for isomorphisms in the category LaxBraidedFunctor C D.

                                          Equations
                                            Instances For

                                              The forgetful functor from lax braided functors to lax monoidal functors.

                                              Equations
                                                Instances For

                                                  The forgetful functor from lax braided functors to lax monoidal functors is fully faithful.

                                                  Equations
                                                    Instances For

                                                      Constructor for isomorphisms between lax braided functors.

                                                      Equations
                                                        Instances For

                                                          A braided functor between braided monoidal categories is a monoidal functor which preserves the braiding.

                                                          Instances
                                                            theorem CategoryTheory.Functor.Braided.ext {C : Type u₁} {inst✝ : Category.{v₁, u₁} C} {inst✝¹ : MonoidalCategory C} {inst✝² : BraidedCategory C} {D : Type u₂} {inst✝³ : Category.{v₂, u₂} D} {inst✝⁴ : MonoidalCategory D} {inst✝⁵ : BraidedCategory D} {F : Functor C D} {x y : F.Braided} (ε : LaxMonoidal.ε F = LaxMonoidal.ε F) (μ : LaxMonoidal.μ F = LaxMonoidal.μ F) (η : OplaxMonoidal.η F = OplaxMonoidal.η F) (δ : OplaxMonoidal.δ F = OplaxMonoidal.δ F) :
                                                            x = y

                                                            A braided category with a faithful braided functor to a symmetric category is itself symmetric.

                                                            Equations
                                                              Instances For
                                                                theorem CategoryTheory.Functor.Braided.toMonoidal_injective {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] {D : Type u₂} [Category.{v₂, u₂} D] [MonoidalCategory D] [BraidedCategory D] (F : Functor C D) :
                                                                Function.Injective (@toMonoidal C inst✝ inst✝¹ inst✝² D inst✝³ inst✝⁴ inst✝⁵ F)

                                                                A multiplicative morphism between commutative monoids gives a braided functor between the corresponding discrete braided monoidal categories.

                                                                Equations
                                                                  def CategoryTheory.MonoidalCategory.tensorμ {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X₁ X₂ Y₁ Y₂ : C) :
                                                                  tensorObj (tensorObj X₁ X₂) (tensorObj Y₁ Y₂) tensorObj (tensorObj X₁ Y₁) (tensorObj X₂ Y₂)

                                                                  Swap the second and third objects in (X₁ ⊗ X₂) ⊗ (Y₁ ⊗ Y₂). This is used to strength the tensor product functor from C × C to C as a monoidal functor.

                                                                  Equations
                                                                    Instances For
                                                                      def CategoryTheory.MonoidalCategory.tensorδ {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X₁ X₂ Y₁ Y₂ : C) :
                                                                      tensorObj (tensorObj X₁ Y₁) (tensorObj X₂ Y₂) tensorObj (tensorObj X₁ X₂) (tensorObj Y₁ Y₂)

                                                                      The inverse of tensorμ.

                                                                      Equations
                                                                        Instances For
                                                                          @[simp]
                                                                          theorem CategoryTheory.MonoidalCategory.tensorμ_tensorδ {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X₁ X₂ Y₁ Y₂ : C) :
                                                                          CategoryStruct.comp (tensorμ X₁ X₂ Y₁ Y₂) (tensorδ X₁ X₂ Y₁ Y₂) = CategoryStruct.id (tensorObj (tensorObj X₁ X₂) (tensorObj Y₁ Y₂))
                                                                          @[simp]
                                                                          theorem CategoryTheory.MonoidalCategory.tensorμ_tensorδ_assoc {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X₁ X₂ Y₁ Y₂ : C) {Z : C} (h : tensorObj (tensorObj X₁ X₂) (tensorObj Y₁ Y₂) Z) :
                                                                          CategoryStruct.comp (tensorμ X₁ X₂ Y₁ Y₂) (CategoryStruct.comp (tensorδ X₁ X₂ Y₁ Y₂) h) = h
                                                                          @[simp]
                                                                          theorem CategoryTheory.MonoidalCategory.tensorδ_tensorμ {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X₁ X₂ Y₁ Y₂ : C) :
                                                                          CategoryStruct.comp (tensorδ X₁ X₂ Y₁ Y₂) (tensorμ X₁ X₂ Y₁ Y₂) = CategoryStruct.id (tensorObj (tensorObj X₁ Y₁) (tensorObj X₂ Y₂))
                                                                          @[simp]
                                                                          theorem CategoryTheory.MonoidalCategory.tensorδ_tensorμ_assoc {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X₁ X₂ Y₁ Y₂ : C) {Z : C} (h : tensorObj (tensorObj X₁ Y₁) (tensorObj X₂ Y₂) Z) :
                                                                          CategoryStruct.comp (tensorδ X₁ X₂ Y₁ Y₂) (CategoryStruct.comp (tensorμ X₁ X₂ Y₁ Y₂) h) = h
                                                                          theorem CategoryTheory.MonoidalCategory.tensorμ_natural {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] {X₁ X₂ Y₁ Y₂ U₁ U₂ V₁ V₂ : C} (f₁ : X₁ Y₁) (f₂ : X₂ Y₂) (g₁ : U₁ V₁) (g₂ : U₂ V₂) :
                                                                          CategoryStruct.comp (tensorHom (tensorHom f₁ f₂) (tensorHom g₁ g₂)) (tensorμ Y₁ Y₂ V₁ V₂) = CategoryStruct.comp (tensorμ X₁ X₂ U₁ U₂) (tensorHom (tensorHom f₁ g₁) (tensorHom f₂ g₂))
                                                                          theorem CategoryTheory.MonoidalCategory.tensorμ_natural_assoc {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] {X₁ X₂ Y₁ Y₂ U₁ U₂ V₁ V₂ : C} (f₁ : X₁ Y₁) (f₂ : X₂ Y₂) (g₁ : U₁ V₁) (g₂ : U₂ V₂) {Z : C} (h : tensorObj (tensorObj Y₁ V₁) (tensorObj Y₂ V₂) Z) :
                                                                          CategoryStruct.comp (tensorHom (tensorHom f₁ f₂) (tensorHom g₁ g₂)) (CategoryStruct.comp (tensorμ Y₁ Y₂ V₁ V₂) h) = CategoryStruct.comp (tensorμ X₁ X₂ U₁ U₂) (CategoryStruct.comp (tensorHom (tensorHom f₁ g₁) (tensorHom f₂ g₂)) h)
                                                                          theorem CategoryTheory.MonoidalCategory.tensorμ_natural_left {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] {X₁ X₂ Y₁ Y₂ : C} (f₁ : X₁ Y₁) (f₂ : X₂ Y₂) (Z₁ Z₂ : C) :
                                                                          CategoryStruct.comp (whiskerRight (tensorHom f₁ f₂) (tensorObj Z₁ Z₂)) (tensorμ Y₁ Y₂ Z₁ Z₂) = CategoryStruct.comp (tensorμ X₁ X₂ Z₁ Z₂) (tensorHom (whiskerRight f₁ Z₁) (whiskerRight f₂ Z₂))
                                                                          theorem CategoryTheory.MonoidalCategory.tensorμ_natural_left_assoc {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] {X₁ X₂ Y₁ Y₂ : C} (f₁ : X₁ Y₁) (f₂ : X₂ Y₂) (Z₁ Z₂ : C) {Z : C} (h : tensorObj (tensorObj Y₁ Z₁) (tensorObj Y₂ Z₂) Z) :
                                                                          CategoryStruct.comp (whiskerRight (tensorHom f₁ f₂) (tensorObj Z₁ Z₂)) (CategoryStruct.comp (tensorμ Y₁ Y₂ Z₁ Z₂) h) = CategoryStruct.comp (tensorμ X₁ X₂ Z₁ Z₂) (CategoryStruct.comp (tensorHom (whiskerRight f₁ Z₁) (whiskerRight f₂ Z₂)) h)
                                                                          theorem CategoryTheory.MonoidalCategory.tensorμ_natural_right {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (Z₁ Z₂ : C) {X₁ X₂ Y₁ Y₂ : C} (f₁ : X₁ Y₁) (f₂ : X₂ Y₂) :
                                                                          CategoryStruct.comp (whiskerLeft (tensorObj Z₁ Z₂) (tensorHom f₁ f₂)) (tensorμ Z₁ Z₂ Y₁ Y₂) = CategoryStruct.comp (tensorμ Z₁ Z₂ X₁ X₂) (tensorHom (whiskerLeft Z₁ f₁) (whiskerLeft Z₂ f₂))
                                                                          theorem CategoryTheory.MonoidalCategory.tensorμ_natural_right_assoc {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (Z₁ Z₂ : C) {X₁ X₂ Y₁ Y₂ : C} (f₁ : X₁ Y₁) (f₂ : X₂ Y₂) {Z : C} (h : tensorObj (tensorObj Z₁ Y₁) (tensorObj Z₂ Y₂) Z) :
                                                                          CategoryStruct.comp (whiskerLeft (tensorObj Z₁ Z₂) (tensorHom f₁ f₂)) (CategoryStruct.comp (tensorμ Z₁ Z₂ Y₁ Y₂) h) = CategoryStruct.comp (tensorμ Z₁ Z₂ X₁ X₂) (CategoryStruct.comp (tensorHom (whiskerLeft Z₁ f₁) (whiskerLeft Z₂ f₂)) h)
                                                                          theorem CategoryTheory.MonoidalCategory.tensor_associativity {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X₁ X₂ Y₁ Y₂ Z₁ Z₂ : C) :
                                                                          CategoryStruct.comp (whiskerRight (tensorμ X₁ X₂ Y₁ Y₂) (tensorObj Z₁ Z₂)) (CategoryStruct.comp (tensorμ (tensorObj X₁ Y₁) (tensorObj X₂ Y₂) Z₁ Z₂) (tensorHom (associator X₁ Y₁ Z₁).hom (associator X₂ Y₂ Z₂).hom)) = CategoryStruct.comp (associator (tensorObj X₁ X₂) (tensorObj Y₁ Y₂) (tensorObj Z₁ Z₂)).hom (CategoryStruct.comp (whiskerLeft (tensorObj X₁ X₂) (tensorμ Y₁ Y₂ Z₁ Z₂)) (tensorμ X₁ X₂ (tensorObj Y₁ Z₁) (tensorObj Y₂ Z₂)))
                                                                          theorem CategoryTheory.MonoidalCategory.tensor_associativity_assoc {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X₁ X₂ Y₁ Y₂ Z₁ Z₂ : C) {Z : C} (h : tensorObj (tensorObj X₁ (tensorObj Y₁ Z₁)) (tensorObj X₂ (tensorObj Y₂ Z₂)) Z) :
                                                                          CategoryStruct.comp (whiskerRight (tensorμ X₁ X₂ Y₁ Y₂) (tensorObj Z₁ Z₂)) (CategoryStruct.comp (tensorμ (tensorObj X₁ Y₁) (tensorObj X₂ Y₂) Z₁ Z₂) (CategoryStruct.comp (tensorHom (associator X₁ Y₁ Z₁).hom (associator X₂ Y₂ Z₂).hom) h)) = CategoryStruct.comp (associator (tensorObj X₁ X₂) (tensorObj Y₁ Y₂) (tensorObj Z₁ Z₂)).hom (CategoryStruct.comp (whiskerLeft (tensorObj X₁ X₂) (tensorμ Y₁ Y₂ Z₁ Z₂)) (CategoryStruct.comp (tensorμ X₁ X₂ (tensorObj Y₁ Z₁) (tensorObj Y₂ Z₂)) h))
                                                                          theorem CategoryTheory.MonoidalCategory.associator_monoidal {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X₁ X₂ X₃ Y₁ Y₂ Y₃ : C) :
                                                                          CategoryStruct.comp (tensorμ (tensorObj X₁ X₂) X₃ (tensorObj Y₁ Y₂) Y₃) (CategoryStruct.comp (whiskerRight (tensorμ X₁ X₂ Y₁ Y₂) (tensorObj X₃ Y₃)) (associator (tensorObj X₁ Y₁) (tensorObj X₂ Y₂) (tensorObj X₃ Y₃)).hom) = CategoryStruct.comp (tensorHom (associator X₁ X₂ X₃).hom (associator Y₁ Y₂ Y₃).hom) (CategoryStruct.comp (tensorμ X₁ (tensorObj X₂ X₃) Y₁ (tensorObj Y₂ Y₃)) (whiskerLeft (tensorObj X₁ Y₁) (tensorμ X₂ X₃ Y₂ Y₃)))
                                                                          theorem CategoryTheory.MonoidalCategory.associator_monoidal_assoc {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X₁ X₂ X₃ Y₁ Y₂ Y₃ : C) {Z : C} (h : tensorObj (tensorObj X₁ Y₁) (tensorObj (tensorObj X₂ Y₂) (tensorObj X₃ Y₃)) Z) :
                                                                          CategoryStruct.comp (tensorμ (tensorObj X₁ X₂) X₃ (tensorObj Y₁ Y₂) Y₃) (CategoryStruct.comp (whiskerRight (tensorμ X₁ X₂ Y₁ Y₂) (tensorObj X₃ Y₃)) (CategoryStruct.comp (associator (tensorObj X₁ Y₁) (tensorObj X₂ Y₂) (tensorObj X₃ Y₃)).hom h)) = CategoryStruct.comp (tensorHom (associator X₁ X₂ X₃).hom (associator Y₁ Y₂ Y₃).hom) (CategoryStruct.comp (tensorμ X₁ (tensorObj X₂ X₃) Y₁ (tensorObj Y₂ Y₃)) (CategoryStruct.comp (whiskerLeft (tensorObj X₁ Y₁) (tensorμ X₂ X₃ Y₂ Y₃)) h))
                                                                          @[simp]
                                                                          theorem CategoryTheory.MonoidalOpposite.mop_braiding {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X Y : C) :
                                                                          (β_ X Y).mop = β_ { unmop := Y } { unmop := X }
                                                                          @[simp]
                                                                          theorem CategoryTheory.MonoidalOpposite.mop_hom_braiding {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X Y : C) :
                                                                          (β_ X Y).hom.mop = (β_ { unmop := Y } { unmop := X }).hom
                                                                          @[simp]
                                                                          theorem CategoryTheory.MonoidalOpposite.mop_inv_braiding {C : Type u₁} [Category.{v₁, u₁} C] [MonoidalCategory C] [BraidedCategory C] (X Y : C) :
                                                                          (β_ X Y).inv.mop = (β_ { unmop := Y } { unmop := X }).inv

                                                                          The identity functor on C, viewed as a functor from C to its monoidal opposite, upgraded to a braided functor.

                                                                          Equations

                                                                            The identity functor on C, viewed as a functor from the monoidal opposite of C to C, upgraded to a braided functor.

                                                                            Equations
                                                                              @[reducible, inline]

                                                                              The braided monoidal category obtained from C by replacing its braiding β_ X Y : X ⊗ Y ≅ Y ⊗ X with the inverse (β_ Y X)⁻¹ : X ⊗ Y ≅ Y ⊗ X. This corresponds to the automorphism of the braid group swapping over-crossings and under-crossings.

                                                                              Equations
                                                                                Instances For

                                                                                  The identity functor from C to C, where the codomain is given the reversed braiding, upgraded to a braided functor.

                                                                                  Equations
                                                                                    Instances For