module Cat.Bi.Base wherePrebicategoriesπ
A (pre)bicategory is the natural higher-dimensional generalisation of a (pre)category. Whereas a precategory has sets, a prebicategory has While this generalisation might seem simple, in reality, we must be very careful when setting up the resulting structure: The key observation is that precategories have no notion of equality of objects, so everything which was an equation in the definition of precategories must be replaced with (sufficiently coherent) specified isomorphisms.
The data of a prebicategory consists of a type of objects and for each a precategory We refer to the type of objects of by and call its inhabitants maps or 1-cells; In the second dimension, between each pair of maps we have a set of 2-cells (sometimes homotopies), written
The composition operation of called vertical composition, will be written As for why itβs called vertical composition, note that it reduces pasting diagrams of the form
record Prebicategory o β β' : Type (lsuc (o β β β β')) where
  no-eta-equality
  field
    Ob  : Type o
    Hom : Ob β Ob β Precategory β β'
  module Hom {A} {B} = Cr (Hom A B)Zooming out to consider the whole bicategory, we see that each object has a specified identity 1-cell as in the case for ordinary categories, but the composition operation, rather than being a function, is a functor. This, intuitively, makes sense: since we have replaced our with we should replace our maps of sets for maps of precategories, i.e., functors.
  field
    id      : β {A} β β Hom A A β
    compose : β {A B C} β Functor (Hom B C ΓαΆ Hom A B) (Hom A C)
  module compose {a} {b} {c} = Fr (compose {a} {b} {c})Before moving on to the isomorphisms witnessing identity and associativity, we introduce a bunch of notation for the different classes of maps and all the composition operations. Observe that the action of the composition functor on homotopies reduces βhorizontalβ pasting diagrams like
whence the name horizontal composition.
  _β¦_ : Ob β Ob β Type β
  A β¦ B = β Hom A B β
  _β_ : β {A B} (f g : A β¦ B) β Type β'
  _β_ {A} {B} f g = Hom.Hom f g
  -- 1-cell composition
  _β_ : β {A B C} (f : B β¦ C) (g : A β¦ B) β A β¦ C
  f β g = compose # (f , g)
  -- vertical 2-cell composition
  _β_ : β {A B} {f g h : A β¦ B} β g β h β f β g β f β h
  _β_ {A} {B} = Hom._β_
  -- horizontal 2-cell composition
  _β_ : β {A B C} {fβ fβ : B β¦ C} (Ξ² : fβ β fβ) {gβ gβ : A β¦ B} (Ξ± : gβ β gβ)
      β (fβ β gβ) β (fβ β gβ)
  _β_ Ξ² Ξ± = compose .Functor.Fβ (Ξ² , Ξ±)
  infixr 30 _β_
  infixr 25 _β_
  infix 35 _β_ _βΆ_
  -- whiskering on the right
  _βΆ_ : β {A B C} (f : B β¦ C) {a b : A β¦ B} (g : a β b) β f β a β f β b
  _βΆ_ {A} {B} {C} f g = compose .Functor.Fβ (Hom.id , g)
  -- whiskering on the left
  _β_ : β {A B C} {a b : B β¦ C} (g : a β b) (f : A β¦ B) β a β f β b β f
  _β_ {A} {B} {C} g f = compose .Functor.Fβ (g , Hom.id)We now move onto the invertible 2-cells witnessing that the chosen identity map is a left- and right- unit element for the composition functor, and that composition is associative. In reality, to get a fully coherent structure, we need these invertible 2-cells to be given as natural isomorphisms, e.g.Β which witnesses that the functor βcompose with the identity 1-cell on the leftβ is naturally isomorphic to the identity functor.
  field
    unitor-l : β {A B} β Cr._β
_ Cat[ Hom A B , Hom A B ] Id (Bi.Right compose id)
    unitor-r : β {A B} β Cr._β
_ Cat[ Hom A B , Hom A B ] Id (Bi.Left compose id)
    associator
      : β {A B C D}
      β Cr._β
_ Cat[ Hom C D ΓαΆ Hom B C ΓαΆ Hom A B , Hom A D ]
        (compose-assocΛ‘ {H = Hom} compose)
        (compose-assocΚ³ {H = Hom} compose)Itβs traditional to refer to the left unitor as to the right unitor as and to the associator as so we set up those abbreviations here too:
  Ξ»β : β {A B} (f : A β¦ B) β id β f β f
  Ξ»β = unitor-l .Cr._β
_.from .Ξ·
  Ξ»β : β {A B} (f : A β¦ B) β f β id β f
  Ξ»β = unitor-l .Cr._β
_.to .Ξ·
  Οβ : β {A B} (f : A β¦ B) β f β id β f
  Οβ = unitor-r .Cr._β
_.from .Ξ·
  Οβ : β {A B} (f : A β¦ B) β f β f β id
  Οβ = unitor-r .Cr._β
_.to .Ξ·
  Οβnat : β {A B} {f f' : A β¦ B} (Ξ² : f β f')
        β Path ((f β id) β f') (Οβ _ β (Ξ² β id)) (Ξ² β Οβ _)
  Οβnat {A} {B} {f} {f'} Ξ² = unitor-r .Cr.from .is-natural f f' Ξ²
  Ξ»βnat : β {A B} {f f' : A β¦ B} (Ξ² : f β f')
        β Path ((id β f) β f') (Ξ»β _ β (id βΆ Ξ²)) (Ξ² β Ξ»β _)
  Ξ»βnat {A} {B} {f} {f'} Ξ² = unitor-l .Cr.from .is-natural f f' Ξ²
  Οβnat : β {A B} {f f' : A β¦ B} (Ξ² : f β f')
        β Path (f β f' β id) (Οβ _ β Ξ²) ((Ξ² β id) β Οβ _)
  Οβnat {A} {B} {f} {f'} Ξ² = unitor-r .Cr.to .is-natural f f' Ξ²
  Ξ»βnat : β {A B} {f f' : A β¦ B} (Ξ² : f β f')
        β Path (f β id β f') (Ξ»β _ β Ξ²) ((id βΆ Ξ²) β Ξ»β _)
  Ξ»βnat {A} {B} {f} {f'} Ξ² = unitor-l .Cr.to .is-natural f f' Ξ²
  Ξ±β : β {A B C D} (f : C β¦ D) (g : B β¦ C) (h : A β¦ B)
     β (f β g) β h β f β (g β h)
  Ξ±β f g h = associator .Cr._β
_.to .Ξ· (f , g , h)
  Ξ±β : β {A B C D} (f : C β¦ D) (g : B β¦ C) (h : A β¦ B)
     β f β (g β h) β (f β g) β h
  Ξ±β f g h = associator .Cr._β
_.from .Ξ· (f , g , h)
  Ξ±βnat : β {A B C D} {f f' : C β¦ D} {g g' : B β¦ C} {h h' : A β¦ B}
        β (Ξ² : f β f') (Ξ³ : g β g') (Ξ΄ : h β h')
        β Path (f β g β h β ((f' β g') β h'))
          (Ξ±β _ _ _ β (Ξ² β (Ξ³ β Ξ΄))) (((Ξ² β Ξ³) β Ξ΄) β Ξ±β _ _ _)
  Ξ±βnat {A} {B} {C} {D} {f} {f'} {g} {g'} {h} {h'} Ξ² Ξ³ Ξ΄ =
    associator .Cr._β
_.from .is-natural (f , g , h) (f' , g' , h') (Ξ² , Ξ³ , Ξ΄)
  Ξ±βnat : β {A B C D} {f f' : C β¦ D} {g g' : B β¦ C} {h h' : A β¦ B}
        β (Ξ² : f β f') (Ξ³ : g β g') (Ξ΄ : h β h')
        β Path ((f β g) β h β (f' β g' β h'))
           (Ξ±β _ _ _ β ((Ξ² β Ξ³) β Ξ΄))
           ((Ξ² β (Ξ³ β Ξ΄)) β Ξ±β _ _ _)
  Ξ±βnat {A} {B} {C} {D} {f} {f'} {g} {g'} {h} {h'} Ξ² Ξ³ Ξ΄ =
    associator .Cr._β
_.to .is-natural (f , g , h) (f' , g' , h') (Ξ² , Ξ³ , Ξ΄)The final data we need are coherences relating the left and right
unitors (the triangle identity, nothing to do with
adjunctions), and one for reducing sequences of associators, the
pentagon identity. As for where the name βpentagonβ
comes from, the path pentagon
witnesses commutativity of the diagram
  field
    triangle
      : β {A B C} (f : B β¦ C) (g : A β¦ B)
      β (Οβ f β g) β Ξ±β f id g β‘ f βΆ Ξ»β g
    pentagon
      : β {A B C D E} (f : D β¦ E) (g : C β¦ D) (h : B β¦ C) (i : A β¦ B)
      β (Ξ±β f g h β i) β Ξ±β f (g β h) i β (f βΆ Ξ±β g h i)
      β‘ Ξ±β (f β g) h i β Ξ±β f g (h β i)Our coherence diagrams for bicategorical data are taken from (Leinster 1998), which contains all the diagrams we have omitted. However, we do not adopt their (dated) terminology of βhomomorphismβ and βstrict homomorphismβ. In contrast with our convention for 1-categories, we refer to bicategories using bold capital letters:
The bicategory of categoriesπ
Just like the prototypal example of categories is the category of sets, the prototypal example of bicategory is the bicategory of categories. We observe that, without a bound of h-level on the spaces of objects (strict categories), the collection of categories of a given universe level does not form a category, but it does form a bicategory.
{-# TERMINATING #-}
Cat : β o β β Prebicategory (lsuc o β lsuc β) (o β β) (o β β)
Cat o β = pb where
  open Prebicategory
  open Functor
  pb : Prebicategory _ _ _
  pb .Ob = Precategory o β
  pb .Hom = Cat[_,_]
  pb .id = IdThe first thing we must compute is that the functor composition operator extends to a functor composition functor, which we have already done (but off-screen, since its construction is very straightforward).
  pb .compose = Fβ-functorThe unitors and associator are almost, but not quite, given by the identity 2-cells, since componentwise the functor composition evaporates, leaving only behind. Unfortunately, this equation is not definitional, so we can not use the identity natural isomorphism directly:
  pb .unitor-r {B = B} = to-natural-iso ni where
    module B = Cr B
    ni : make-natural-iso {D = Cat[ _ , _ ]} _ _
    ni .make-natural-iso.eta x = NT (Ξ» _ β B.id) Ξ» _ _ _ β B.id-comm-sym
    ni .make-natural-iso.inv x = NT (Ξ» _ β B.id) Ξ» _ _ _ β B.id-comm-sym
    ni .make-natural-iso.etaβinv x = ext Ξ» _ β B.idl _
    ni .make-natural-iso.invβeta x = ext Ξ» _ β B.idl _
    ni .make-natural-iso.natural x y f =
      ext Ξ» _ β B.idr _ β ap (B._β _) (y .F-id)
  pb .unitor-l {B = B} = to-natural-iso ni where
    module B = Cr B
    ni : make-natural-iso {D = Cat[ _ , _ ]} _ _
    ni .make-natural-iso.eta x = NT (Ξ» _ β B.id) Ξ» _ _ _ β B.id-comm-sym
    ni .make-natural-iso.inv x = NT (Ξ» _ β B.id) Ξ» _ _ _ β B.id-comm-sym
    ni .make-natural-iso.etaβinv x = ext Ξ» _ β B.idl _
    ni .make-natural-iso.invβeta x = ext Ξ» _ β B.idl _
    ni .make-natural-iso.natural x y f = ext Ξ» _ β B.idr _ β B.id-comm
  pb .associator {A} {B} {C} {D} = to-natural-iso ni where
    module D = Cr D
    ni : make-natural-iso {D = Cat[ _ , _ ]} _ _
    ni .make-natural-iso.eta x = NT (Ξ» _ β D.id) Ξ» _ _ _ β D.id-comm-sym
    ni .make-natural-iso.inv x = NT (Ξ» _ β D.id) Ξ» _ _ _ β D.id-comm-sym
    ni .make-natural-iso.etaβinv x = ext Ξ» _ β D.idl _
    ni .make-natural-iso.invβeta x = ext Ξ» _ β D.idl _
    ni .make-natural-iso.natural x y f = ext Ξ» _ β
      D.idr _ Β·Β· D.pushl (y .fst .F-β _ _) Β·Β· D.introl refl
  pb .triangle {C = C} f g = ext Ξ» _ β Cr.idr C _
  pb .pentagon {E = E} f g h i = ext Ξ» _ β apβ E._β_
    (E.eliml (ap (f .Fβ) (ap (g .Fβ) (h .F-id)) Β·Β· ap (f .Fβ) (g .F-id) Β·Β· f .F-id))
    (E.elimr (E.eliml (f .F-id)))
    where module E = Cr ELax functorsπ
In the same way that the definition of bicategory is obtained by starting with the definition of category and replacing the by (and adding coherence data to make sure the resulting structure is well-behaved), one can start with the definition of functor and replace the function between by functors between
However, when talking about general bicategories, we are faced with a choice: We could generalise the functoriality axioms to natural isomorphisms, keeping with the fact that equations are invertible, but we could also drop this invertibility requirement, and work only with natural transformations When these are not invertible, the resulting structure is called a lax functor; When they are, we talk about pseudofunctors instead.
record
  Lax-functor (B : Prebicategory o β β') (C : Prebicategory oβ ββ ββ')
    : Type (o β oβ β β β ββ β β' β ββ') where
  private
    module B = Prebicategory B
    module C = Prebicategory C
  field
    Pβ : B.Ob β C.Ob
    Pβ : β {A B} β Functor (B.Hom A B) (C.Hom (Pβ A) (Pβ B))The resulting structure has βdirected functorialityβ, witnessed by
the compositor and unitor natural transformations, which
have components
and
    compositor
      : β {A B C}
      β C.compose Fβ (Pβ {B} {C} FΓ Pβ {A} {B}) => Pβ Fβ B.compose
    unitor : β {A} β C.id C.β Pβ .Functor.Fβ (B.id {A = A})Additionally, we require the following three equations to hold,
relating the compositor transformation to the associators, and the three
unitors between themselves. We sketch the diagram which hexagon witnesses commutativity for, but
leave the right-unit and left-unit diagrams undrawn (theyβre
boring commutative squares).
  field
    hexagon
      : β {a b c d} (f : c B.β¦ d) (g : b B.β¦ c) (h : a B.β¦ b)
      β β (B.Ξ±β f g h) C.β Ξ³β (f B.β g) h C.β (Ξ³β f g C.β β h)
      β‘ Ξ³β f (g B.β h) C.β (β f C.βΆ Ξ³β g h) C.β C.Ξ±β (β f) (β g) (β h)
    right-unit
      : β {a b} (f : a B.β¦ b)
      β β (B.Οβ f) C.β Ξ³β f B.id C.β (β f C.βΆ unitor) β‘ C.Οβ (β f)
    left-unit
      : β {a b} (f : a B.β¦ b)
      β β (B.Ξ»β f) C.β Ξ³β B.id f C.β (unitor C.β β f) β‘ C.Ξ»β (β f)Pseudofunctorsπ
As mentioned above, a lax functor with invertible unitors and
compositor is called a pseudofunctor. Every
pseudofunctor has an underlying lax
functor. Since invertibility is a property of 2-cells (rather
than structure on 2-cells), βbeing pseudoβ is a property of lax
functors, not additional structure on lax functors.
record
  Pseudofunctor (B : Prebicategory o β β') (C : Prebicategory oβ ββ ββ')
    : Type (o β oβ β β β ββ β β' β ββ') where
  private
    module B = Prebicategory B
    module C = Prebicategory C
  field
    lax : Lax-functor B C
  open Lax-functor lax public
  field
    unitor-inv
      : β {a} β C.Hom.is-invertible (unitor {a})
    compositor-inv
      : β {a b c} (f : b B.β¦ c) (g : a B.β¦ b) β C.Hom.is-invertible (Ξ³β f g)
  Ξ³β : β {a b c} (f : b B.β¦ c) (g : a B.β¦ b)
    β β (f B.β g) C.β β f C.β β g
  Ξ³β f g = compositor-inv f g .Cr.is-invertible.inv
  Ο
β : β {a} β β B.id C.β C.id
  Ο
β {a} = unitor-inv {a = a} .Cr.is-invertible.invLax transformationsπ
By dropping the invertibility requirement when generalising natural transformations to lax functors, we obtain the type of lax transformations between lax functors. If every 2-cell component of the lax transformation is invertible, we refer to it as a pseudonatural transformation. We omit the word βnaturalβ in βlax natural transformationβ for brevity.
The transformation which witnesses directed naturality for a lax
transformation is called the naturator. In components, it witnesses
commutativity of the diagram
and thus consists of a natural family of 2-cells
  record Lax-transfor : Type (o β β β ββ β β' β ββ') where
    field
      Ο : β A β F.β A C.β¦ G.β A
      naturator
        : β {a b}
        β preaction C (Ο b) Fβ G.Pβ => postaction C (Ο a) Fβ F.Pβ
    Ξ½β : β {a b} (f : a B.β¦ b) β G.β f C.β Ο a C.β Ο b C.β F.β f
    Ξ½β = naturator .Ξ·The naturator is required to be compatible with the compositor and unitor natural transformations of its source and target functors, which boil down to commutativity of the nightmarish diagrams in (Leinster 1998, sec. 1.2).
    field
      Ξ½-compositor
        : β {a b c} (f : b B.β¦ c) (g : a B.β¦ b)
        β Ξ½β (f B.β g) C.β (G.Ξ³β f g C.β Ο a)
        β‘   (Ο c C.βΆ F.Ξ³β f g)
        C.β C.Ξ±β (Ο c) (F.β f) (F.β g)
        C.β (Ξ½β f C.β F.β g)
        C.β C.Ξ±β (G.β f) (Ο b) (F.β g)
        C.β (G.β f C.βΆ Ξ½β g)
        C.β C.Ξ±β (G.β f) (G.β g) (Ο a)
      Ξ½-unitor
        : β {a}
        β Ξ½β (B.id {a}) C.β (G.unitor C.β Ο a)
        β‘ (Ο a C.βΆ F.unitor) C.β C.Οβ (Ο a) C.β C.Ξ»β (Ο a)A lax transformation with invertible naturator is called a pseudonatural transformation.
  record Pseudonatural : Type (o β β β ββ β β' β ββ') where
    field
      lax : Lax-transfor
    open Lax-transfor lax public
    field
      naturator-inv : β {a b} (f : a B.β¦ b) β C.Hom.is-invertible (Ξ½β f)
    Ξ½β : β {a b} (f : a B.β¦ b) β Ο b C.β F.β f C.β G.β f C.β Ο a
    Ξ½β f = naturator-inv f .Cr.is-invertible.invWe abbreviate the types of lax- and pseudonatural transformations by
_=>β_
and _=>β_,
respectively.
  _=>β_ = Lax-transfor
  _=>β_ = PseudonaturalModificationsπ
When dealing with 1-categorical data (categories, functors, and natural transformations), the commutativity in 2-cells is witnessed by equations in a set, which are trivial. When talking about bicategorical data, however, the naturality of a lax transformation is witnessed by a family of non-trivial 2-cells. Therefore, it is fruitful to consider transformations which affect this data: a natural family of 2-cells. This is called a modification between lax (or pseudo) transformations. Since we are directly dealing with sets (the sets of 2-cells), modifications are the simplest bicategorical widget to define.
  record Modification : Type (o β β β ββ') where
    field
      Ξ : β a β Ο.Ο a C.β Ο'.Ο a
      is-natural
        : β {a b} {f : a B.β¦ b}
        β Ο'.Ξ½β f C.β (G.β f C.βΆ Ξ a)
        β‘ (Ξ b C.β F.β f) C.β Ο.Ξ½β fIn a diagram, we display a modification as a 3-cell, i.e., a morphism (modification) between morphisms (lax transformations) between morphisms (lax functors) between objects (bicategories), and accordingly draw them with super-heavy arrows, as in the diagram below. Fortunately we will not often stumble onto diagrams of bicategories, rather studying diagrams in bicategories, which are (mercifully) limited to 2-cells.
References
- Leinster, Tom. 1998. βBasic Bicategories.β arXiv. https://doi.org/10.48550/ARXIV.MATH/9810017.