Opposite bicategories #
We construct the 1-cell opposite of a bicategory B, called Bᵒᵖ. It is defined as follows
- The objects of
Bᵒᵖcorrespond to objects ofB. - The morphisms
X ⟶ YinBᵒᵖare the morphismsY ⟶ XinB. - The 2-morphisms
f ⟶ ginBᵒᵖare the 2-morphismsf ⟶ ginB. In other words, the directions of the 2-morphisms are preserved.
Remarks #
There are multiple notions of opposite categories for bicategories.
- There is 1-cell dual
Bᵒᵖas defined above. - There is the 2-cell dual,
Cᶜᵒwhere only the natural transformations are reversed - There is the bi-dual
Cᶜᵒᵒᵖwhere the directions of both the morphisms and the natural transformations are reversed.
TODO #
- Define the 2-cell dual
Cᶜᵒ. - Provide various lemmas for going between
LocallyDiscrete Cᵒᵖand(LocallyDiscrete C)ᵒᵖ.
Note: Cᶜᵒᵒᵖ is WIP by Christian Merten.
Type synonym for 2-morphisms in the opposite bicategory.
- op2' :: (
Bᵒᵖpreserves the direction of all 2-morphisms inB- )
Instances For
Equations
- One or more equations did not get rendered due to their size.
Synonym for constructor of Hom2 where the 1-morphisms f and g lie in B and not Bᵒᵖ.
Equations
- Bicategory.Opposite.op2 η = { unop2 := η }
Instances For
The natural functor from the hom-category a ⟶ b in B to its bicategorical opposite
bop b ⟶ bop a.
Equations
- Bicategory.Opposite.opFunctor a b = { obj := fun (f : a ⟶ b) => f.op, map := fun {X Y : a ⟶ b} (η : X ⟶ Y) => Bicategory.Opposite.op2 η, map_id := ⋯, map_comp := ⋯ }
Instances For
The functor from the hom-category a ⟶ b in Bᵒᵖ to its bicategorical opposite
unop b ⟶ unop a.
Equations
Instances For
A 2-isomorphism in B gives a 2-isomorphism in Bᵒᵖ
Equations
- η.op2 = (Bicategory.Opposite.opFunctor a b).mapIso η
Instances For
A 2-isomorphism in B gives a 2-isomorphism in Bᵒᵖ
Equations
- η.op2_unop = (Bicategory.Opposite.opFunctor (Opposite.unop b) (Opposite.unop a)).mapIso η
Instances For
A 2-isomorphism in Bᵒᵖ gives a 2-isomorphism in B
Equations
- η.unop2 = (Bicategory.Opposite.unopFunctor a b).mapIso η
Instances For
A 2-isomorphism in Bᵒᵖ gives a 2-isomorphism in B
Equations
- η.unop2_op = (Bicategory.Opposite.unopFunctor (Opposite.op b) (Opposite.op a)).mapIso η
Instances For
The 1-cell dual bicategory Bᵒᵖ.
It is defined as follows.
- The objects of
Bᵒᵖcorrespond to objects ofB. - The morphisms
X ⟶ YinBᵒᵖare the morphismsY ⟶ XinB. - The 2-morphisms
f ⟶ ginBᵒᵖare the 2-morphismsf ⟶ ginB. In other words, the directions of the 2-morphisms are preserved.
Equations
- One or more equations did not get rendered due to their size.