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
tensorObj : C → C → C
tensorHom : (X₁ ⟶ Y₁) → (X₂ ⟶ Y₂) → ((X₁ ⊗ X₂) ⟶ (Y₁ ⊗ Y₂))
and allow use of the overloaded notation⊗
for both. The unitors and associator are provided componentwise.
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:
whiskerLeft (X : C) {Y₁ Y₂ : C} (f : Y₁ ⟶ Y₂) : X ⊗ Y₁ ⟶ X ⊗ Y₂
, denoted byX ◁ f
,whiskerRight {X₁ X₂ : C} (f : X₁ ⟶ X₂) (Y : C) : X₁ ⊗ Y ⟶ X₂ ⊗ Y
, denoted byf ▷ Y
. These are products of an object and a morphism (the terminology "whiskering" is borrowed from 2-category theory). The tensor product of morphismstensorHom
can be defined in terms of the whiskerings. There are two possible such definitions, which are related by the exchange property of the whiskerings. These two definitions are accessed bytensorHom_def
andtensorHom_def'
. By default,tensorHom
is defined so thattensorHom_def
holds definitionally.
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.
References #
- Tensor categories, Etingof, Gelaki, Nikshych, Ostrik, http://www-math.mit.edu/~etingof/egnobookfinal.pdf
- https://stacks.math.columbia.edu/tag/0FFK.
- tensorObj : C → C → C
curried tensor product of objects
- whiskerLeft : (X : C) → {Y₁ Y₂ : C} → (Y₁ ⟶ Y₂) → (CategoryTheory.MonoidalCategory.tensorObj X Y₁ ⟶ CategoryTheory.MonoidalCategory.tensorObj X Y₂)
left whiskering for morphisms
- whiskerRight : {X₁ X₂ : C} → (X₁ ⟶ X₂) → (Y : C) → CategoryTheory.MonoidalCategory.tensorObj X₁ Y ⟶ CategoryTheory.MonoidalCategory.tensorObj X₂ Y
right whiskering for morphisms
- tensorHom : {X₁ Y₁ X₂ Y₂ : C} → (X₁ ⟶ Y₁) → (X₂ ⟶ Y₂) → (CategoryTheory.MonoidalCategory.tensorObj X₁ X₂ ⟶ CategoryTheory.MonoidalCategory.tensorObj Y₁ Y₂)
Tensor product of identity maps is the identity:
(𝟙 X₁ ⊗ 𝟙 X₂) = 𝟙 (X₁ ⊗ X₂)
- tensorHom_def : ∀ {X₁ Y₁ X₂ Y₂ : C} (f : X₁ ⟶ Y₁) (g : X₂ ⟶ Y₂), CategoryTheory.MonoidalCategory.tensorHom f g = CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.whiskerRight f X₂) (CategoryTheory.MonoidalCategory.whiskerLeft Y₁ g)
- tensor_id : ∀ (X₁ X₂ : C), CategoryTheory.MonoidalCategory.tensorHom (CategoryTheory.CategoryStruct.id X₁) (CategoryTheory.CategoryStruct.id X₂) = CategoryTheory.CategoryStruct.id (CategoryTheory.MonoidalCategory.tensorObj X₁ X₂)
Tensor product of identity maps is the identity:
(𝟙 X₁ ⊗ 𝟙 X₂) = 𝟙 (X₁ ⊗ X₂)
- tensor_comp : ∀ {X₁ Y₁ Z₁ X₂ Y₂ Z₂ : C} (f₁ : X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) (g₁ : Y₁ ⟶ Z₁) (g₂ : Y₂ ⟶ Z₂), CategoryTheory.MonoidalCategory.tensorHom (CategoryTheory.CategoryStruct.comp f₁ g₁) (CategoryTheory.CategoryStruct.comp f₂ g₂) = CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.tensorHom f₁ f₂) (CategoryTheory.MonoidalCategory.tensorHom g₁ g₂)
Composition of tensor products is tensor product of compositions:
(f₁ ⊗ g₁) ∘ (f₂ ⊗ g₂) = (f₁ ∘ f₂) ⊗ (g₁ ⊗ g₂)
- tensorUnit' : C
The tensor unity in the monoidal structure
𝟙_ C
- associator : (X Y Z : C) → CategoryTheory.MonoidalCategory.tensorObj (CategoryTheory.MonoidalCategory.tensorObj X Y) Z ≅ CategoryTheory.MonoidalCategory.tensorObj X (CategoryTheory.MonoidalCategory.tensorObj Y Z)
The associator isomorphism
(X ⊗ Y) ⊗ Z ≃ X ⊗ (Y ⊗ Z)
- whiskerLeft_id : ∀ (X Y : C), CategoryTheory.MonoidalCategory.whiskerLeft X (CategoryTheory.CategoryStruct.id Y) = CategoryTheory.CategoryStruct.id (CategoryTheory.MonoidalCategory.tensorObj X Y)
- id_whiskerRight : ∀ (X Y : C), CategoryTheory.MonoidalCategory.whiskerRight (CategoryTheory.CategoryStruct.id X) Y = CategoryTheory.CategoryStruct.id (CategoryTheory.MonoidalCategory.tensorObj X Y)
- associator_naturality : ∀ {X₁ X₂ X₃ Y₁ Y₂ Y₃ : C} (f₁ : X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) (f₃ : X₃ ⟶ Y₃), CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.tensorHom (CategoryTheory.MonoidalCategory.tensorHom f₁ f₂) f₃) (CategoryTheory.MonoidalCategory.associator Y₁ Y₂ Y₃).hom = CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.associator X₁ X₂ X₃).hom (CategoryTheory.MonoidalCategory.tensorHom f₁ (CategoryTheory.MonoidalCategory.tensorHom f₂ f₃))
Naturality of the associator isomorphism:
(f₁ ⊗ f₂) ⊗ f₃ ≃ f₁ ⊗ (f₂ ⊗ f₃)
- leftUnitor : (X : C) → CategoryTheory.MonoidalCategory.tensorObj CategoryTheory.MonoidalCategory.tensorUnit' X ≅ X
The left unitor:
𝟙_ C ⊗ X ≃ X
- leftUnitor_naturality : ∀ {X Y : C} (f : X ⟶ Y), CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.tensorHom (CategoryTheory.CategoryStruct.id CategoryTheory.MonoidalCategory.tensorUnit') f) (CategoryTheory.MonoidalCategory.leftUnitor Y).hom = CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.leftUnitor X).hom f
Naturality of the left unitor, commutativity of
𝟙_ C ⊗ X ⟶ 𝟙_ C ⊗ Y ⟶ Y
and𝟙_ C ⊗ X ⟶ X ⟶ Y
- rightUnitor : (X : C) → CategoryTheory.MonoidalCategory.tensorObj X CategoryTheory.MonoidalCategory.tensorUnit' ≅ X
The right unitor:
X ⊗ 𝟙_ C ≃ X
- rightUnitor_naturality : ∀ {X Y : C} (f : X ⟶ Y), CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.tensorHom f (CategoryTheory.CategoryStruct.id CategoryTheory.MonoidalCategory.tensorUnit')) (CategoryTheory.MonoidalCategory.rightUnitor Y).hom = CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.rightUnitor X).hom f
Naturality of the right unitor: commutativity of
X ⊗ 𝟙_ C ⟶ Y ⊗ 𝟙_ C ⟶ Y
andX ⊗ 𝟙_ C ⟶ X ⟶ Y
- pentagon : ∀ (W X Y Z : C), CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.tensorHom (CategoryTheory.MonoidalCategory.associator W X Y).hom (CategoryTheory.CategoryStruct.id Z)) (CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.associator W (CategoryTheory.MonoidalCategory.tensorObj X Y) Z).hom (CategoryTheory.MonoidalCategory.tensorHom (CategoryTheory.CategoryStruct.id W) (CategoryTheory.MonoidalCategory.associator X Y Z).hom)) = CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.associator (CategoryTheory.MonoidalCategory.tensorObj W X) Y Z).hom (CategoryTheory.MonoidalCategory.associator W X (CategoryTheory.MonoidalCategory.tensorObj Y Z)).hom
The pentagon identity relating the isomorphism between
X ⊗ (Y ⊗ (Z ⊗ W))
and((X ⊗ Y) ⊗ Z) ⊗ W
- triangle : ∀ (X Y : C), CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.associator X CategoryTheory.MonoidalCategory.tensorUnit' Y).hom (CategoryTheory.MonoidalCategory.tensorHom (CategoryTheory.CategoryStruct.id X) (CategoryTheory.MonoidalCategory.leftUnitor Y).hom) = CategoryTheory.MonoidalCategory.tensorHom (CategoryTheory.MonoidalCategory.rightUnitor X).hom (CategoryTheory.CategoryStruct.id Y)
The identity relating the isomorphisms between
X ⊗ (𝟙_C ⊗ Y)
,(X ⊗ 𝟙_C) ⊗ Y
andX ⊗ Y
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.
See https://stacks.math.columbia.edu/tag/0FFK.
Instances
The tensor unity in the monoidal structure 𝟙_ C
Equations
- CategoryTheory.MonoidalCategory.tensorUnit C = CategoryTheory.MonoidalCategory.tensorUnit'
Instances For
Notation for tensorObj
, the tensor product of objects in a monoidal category
Equations
- One or more equations did not get rendered due to their size.
Instances For
Notation for the whiskerLeft
operator of monoidal categories
Equations
- One or more equations did not get rendered due to their size.
Instances For
Notation for the whiskerRight
operator of monoidal categories
Equations
- One or more equations did not get rendered due to their size.
Instances For
Notation for tensorHom
, the tensor product of morphisms in a monoidal category
Equations
- One or more equations did not get rendered due to their size.
Instances For
Notation for tensorUnit
, the two-sided identity of ⊗
Equations
- CategoryTheory.MonoidalCategory.«term𝟙_» = Lean.ParserDescr.node `CategoryTheory.MonoidalCategory.term𝟙_ 1024 (Lean.ParserDescr.symbol "𝟙_")
Instances For
Notation for the monoidal associator
: (X ⊗ Y) ⊗ Z) ≃ X ⊗ (Y ⊗ Z)
Equations
- CategoryTheory.MonoidalCategory.termα_ = Lean.ParserDescr.node `CategoryTheory.MonoidalCategory.termα_ 1024 (Lean.ParserDescr.symbol "α_")
Instances For
Notation for the leftUnitor
: 𝟙_C ⊗ X ≃ X
Equations
- CategoryTheory.MonoidalCategory.«termλ_» = Lean.ParserDescr.node `CategoryTheory.MonoidalCategory.termλ_ 1024 (Lean.ParserDescr.symbol "λ_")
Instances For
Notation for the rightUnitor
: X ⊗ 𝟙_C ≃ X
Equations
- CategoryTheory.MonoidalCategory.termρ_ = Lean.ParserDescr.node `CategoryTheory.MonoidalCategory.termρ_ 1024 (Lean.ParserDescr.symbol "ρ_")
Instances For
The tensor product of two isomorphisms is an isomorphism.
Equations
- f ⊗ g = CategoryTheory.Iso.mk (CategoryTheory.MonoidalCategory.tensorHom f.hom g.hom) (CategoryTheory.MonoidalCategory.tensorHom f.inv g.inv)
Instances For
Notation for tensorIso
, the tensor product of isomorphisms
Equations
- CategoryTheory.«term_⊗_» = Lean.ParserDescr.trailingNode `CategoryTheory.term_⊗_ 70 71 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ⊗ ") (Lean.ParserDescr.cat `term 70))
Instances For
The lemmas in the next section are true by coherence, but we prove them directly as they are used in proving the coherence theorem.
A constructor for monoidal categories that requires tensorHom
instead of whiskerLeft
and
whiskerRight
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The tensor product expressed as a functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The left-associated triple tensor product as a functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The right-associated triple tensor product as a functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The functor fun X ↦ 𝟙_ C ⊗ X
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The functor fun X ↦ X ⊗ 𝟙_ C
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The associator as a natural isomorphism.
Equations
- CategoryTheory.MonoidalCategory.associatorNatIso C = CategoryTheory.NatIso.ofComponents fun X => CategoryTheory.MonoidalCategory.associator X.fst X.snd.fst X.snd.snd
Instances For
The left unitor as a natural isomorphism.
Equations
Instances For
The right unitor as a natural isomorphism.
Equations
Instances For
Tensoring on the left with a fixed object, as a functor.
Equations
- One or more equations did not get rendered due to their size.
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
Tensoring on the right with a fixed object, as a functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Tensoring on the left, as a functor from C
into endofunctors of C
.
TODO: show this is an op-monoidal functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Tensoring on the right, as a functor from C
into endofunctors of C
.
We later show this is a monoidal functor.
Equations
- One or more equations did not get rendered due to their size.
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
Equations
- One or more equations did not get rendered due to their size.