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.
The bifunctor (F -) ⊗ -
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The bifunctor - ⊗ (F -)
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The bifunctor F - ⊗ F -
.
Equations
Instances For
The bifunctor F (- ⊗ -)
.
Equations
Instances For
The trifunctor (F - ⊗ F -) ⊗ F -
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The trifunctor F - ⊗ (F - ⊗ F -)
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The trifunctor F (- ⊗ -) ⊗ F -
.
Equations
Instances For
The trifunctor F - ⊗ F (- ⊗ -)
.
Equations
Instances For
The trifunctor F ((- ⊗ -) ⊗ -)
Equations
Instances For
The trifunctor F (- ⊗ (- ⊗ -))
Equations
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 top left map in the associativity hexagon.
Equations
Instances For
The middle left map in the associativity hexagon.
Equations
Instances For
The bottom left map in the associativity hexagon.
Equations
Instances For
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
The middle right map in the associativity hexagon.
Equations
Instances For
The bottom right map in the associativity hexagon.
Equations
Instances For
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
The left map in the left unitality square.
Equations
Instances For
The top map in the left unitality square.
Equations
Instances For
The bottom map in the left unitality square.
Equations
Instances For
The left map in the right unitality square.
Equations
Instances For
The top map in the right unitality square.
Equations
Instances For
The bottom map in the right unitality square.
Equations
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 top left map in the oplax associativity hexagon.
Equations
Instances For
The middle left map in the oplax associativity hexagon.
Equations
Instances For
The bottom left map 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 left maps in the oplax associativity hexagon.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The top right map in the oplax associativity hexagon.
Equations
Instances For
The middle right map in the oplax associativity hexagon.
Equations
Instances For
The bottom right map in the oplax associativity hexagon.
Equations
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
The left map in the oplax left unitality square.
Equations
Instances For
The top map in the oplax left unitality square.
Equations
Instances For
The bottom map in the oplax left unitality square.
Equations
Instances For
The left map in the oplax right unitality square.
Equations
Instances For
The top map in the oplax right unitality square.
Equations
Instances For
The bottom map in the oplax right unitality square.
Equations
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
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.