Theory Coherence
section "Coherence"
theory Coherence
imports Bicategory
begin
text ‹
\sloppypar
In this section, we generalize to bicategories the proof of the Coherence Theorem
that we previously gave for monoidal categories
(see ‹MonoidalCategory.evaluation_map.coherence› in @{session MonoidalCategory}).
As was the case for the previous proof, the current proof takes a syntactic approach.
First we define a formal ``bicategorical language'' of terms constructed using
syntactic operators that correspond to the various operations (vertical and horizontal
composition, associators and unitors) found in a bicategory.
Terms of the language are classified as formal ``arrows'', ``identities'', or ``objects''
according to the syntactic operators used in their formation.
A class of terms called ``canonical'' is also defined in this way.
Functions that map ``arrows'' to their ``domain'' and ``codomain'', and to their
``source'' and ``target'' are defined by recursion on the structure of terms.
Next, we define a notion of ``normal form'' for terms in this language and we
give a recursive definition of a function that maps terms to their normal forms.
Normalization moves vertical composition inside of horizontal composition and
``flattens'' horizontal composition by associating all horizontal compositions to the right.
In addition, normalization deletes from a term any horizontal composites involving an arrow
and its source or target, replacing such composites by just the arrow itself.
We then define a ``reduction function'' that maps each identity term ‹t› to a
``canonical'' term ‹t❙↓› that connects ‹t› with its normal form. The definition of reduction
is also recursive, but it is somewhat more complex than normalization in that it
involves two mutually recursive functions: one that applies to any identity term
and another that applies only to terms that are the horizontal composite
of two identity terms.
The next step is to define an ``evaluation function'' that evaluates terms in a given
bicategory (which is left as an unspecified parameter). We show that evaluation respects
bicategorical structure:
the domain, codomain, source, and target mappings on terms correspond under evaluation
to the actual domain, codomain, source and target mappings on the given bicategory,
the vertical and horizontal composition on terms correspond to the actual vertical
and horizontal composition of the bicategory, and unit and associativity terms evaluate
to the actual unit and associativity isomorphisms of the bicategory.
In addition, ``object terms'' evaluate to objects (\emph{i.e.}~0-cells),
``identity terms'' evaluate to identities (\emph{i.e.}~1-cells),
``arrow terms'' evaluate to arrows (\emph{i.e.}~2-cells), and ``canonical terms'' evaluate
to canonical isomorphisms.
A term is defined to be ``coherent'' if, roughly speaking, it is a formal arrow
whose evaluation commutes with the evaluations of the reductions to normal form of
its domain and codomain.
We then prove the Coherence Theorem, expressed in the form: ``every arrow is coherent.''
This implies a more classical version of the Coherence Theorem, which says that:
``syntactically parallel arrows with the same normal form have equal evaluations''.
›
subsection "Bicategorical Language"
text ‹
For the most part, the definition of the ``bicategorical language'' of terms is
a straightforward generalization of the ``monoidal language'' that we used for
monoidal categories.
Some modifications are required, however, due to the fact that horizontal composition
in a bicategory is a partial operation, whereas the the tensor product in a monoidal
category is well-defined for all pairs of arrows.
One difference is that we have found it necessary to introduce a new class of primitive
terms whose elements represent ``formal objects'', so that there is some way to
identify the source and target of what would otherwise be an empty horizontal composite.
This was not an issue for monoidal categories, because the totality of horizontal
composition meant that there was no need for syntactically defined sources and targets.
Another difference is what we have chosen for the ``generators'' of the language
and how they are used to form primitive terms. For monoidal categories,
we supposed that we were given a category ‹C› and the syntax contained a constructor
to form a primitive term corresponding to each arrow of ‹C›.
We assumed a category as the given data, rather than something less structured,
such as a graph, because we were primarily interested in the tensor product and
the associators and unitors, and were relatively uninterested in the strictly
associative and unital composition of the underlying category.
For bicategories, we also take the vertical composition as given for the same
reasons; however, this is not yet sufficient due to the fact that horizontal
composition in a bicategory is a partial operation, in contrast to the tensor
product in a monoidal category, which is defined for all pairs of arrows.
To deal with this issue, for bicategories we assume that source and target
mappings are also given, so that the given data forms a category with
``horizontal homs''. The given source and target mappings are extended to all terms
and used to define when two terms are ``formally horizontally composable''.
›
locale bicategorical_language =
category V +
horizontal_homs V src trg
for V :: "'a comp" (infixr "⋅" 55)
and src :: "'a ⇒ 'a"
and trg :: "'a ⇒ 'a"
begin
text ‹
Constructor ‹Prim⇩0› is used to construct ``formal objects'' and ‹Prim› is used to
construct primitive terms that are not formal objects.
›