module 1Lab.Path.Groupoid where
Types are groupoidsπ
The Path
types equip every Type
with the structure of an
-groupoid.
The higher structure of a type begins with its inhabitants (the
0-cells); Then, there are the paths between inhabitants - these are
inhabitants of the type Path A x y
, which are the 1-cells
in A
. Then, we can consider the inhabitants of
Path (Path A x y) p q
, which are homotopies
between paths.
This construction iterates forever - and, at every stage, we have
that the
-cells
in A
are the 0-cells of some other type (e.g.: The 1-cells
of A
are the 0-cells of Path A ...
).
Furthermore, this structure is weak: The laws,
e.g.Β associativity of composition, only hold up to a homotopy. These
laws satisfy their own laws β again using associativity as an example,
the associator satisfies the pentagon identity.
These laws can be proved in two ways: Using path induction, or directly with a cubical argument. Here, we do both.
Book-styleπ
This is the approach taken in the HoTT book. We fix a type, and some variables of that type, and some paths between variables of that type, so that each definition doesnβt start with 12 parameters.
module WithJ where
private variable
: Level
β : Type β
A : A w x y z
First, we (re)define the operations using J
. These will be closer to the structure
given in the book.
_β_ : x β‘ y β y β‘ z β x β‘ z
_β_ {x = x} {y} {z} = J (Ξ» y _ β y β‘ z β x β‘ z) (Ξ» x β x)
First we define path composition. Then, we can prove that the
identity path - refl
- acts as an
identity for path composition.
: (p : x β‘ y) β p β refl β‘ p
β-idr {x = x} {y = y} p =
β-idr (Ξ» _ p β p β refl β‘ p)
J (happly (J-refl (Ξ» y _ β y β‘ y β x β‘ y) (Ξ» x β x)) _)
p
This isnβt as straightforward as it would be in βBook HoTTβ because -
remember - J
doesnβt compute
definitionally, only up to the path J-refl
. Now the other identity law:
: (p : y β‘ z) β refl β p β‘ p
β-idl {y = y} {z = z} p = happly (J-refl (Ξ» y _ β y β‘ z β y β‘ z) (Ξ» x β x)) p β-idl
This case we get for less since itβs essentially the computation rule
for J
.
: (p : w β‘ x) (q : x β‘ y) (r : y β‘ z)
β-assoc β p β (q β r) β‘ (p β q) β r
{w = w} {x = x} {y = y} {z = z} p q r =
β-assoc (Ξ» x p β (q : x β‘ y) (r : y β‘ z) β p β (q β r) β‘ (p β q) β r) lemma p q r
J where
: (q : w β‘ y) (r : y β‘ z)
lemma β (refl β (q β r)) β‘ ((refl β q) β r)
=
lemma q r (refl β (q β r)) β‘β¨ β-idl (q β r) β©
(ap (Ξ» e β e β r) (β-idl q)) β©
q β r β‘β¨ sym (refl β q) β r β
The associativity rule is trickier to prove, since we do inductive
where the motive is a dependent product. What weβre doing can be
summarised using words: By induction, it suffices to assume
p
is refl. Then, what we want to show is
(refl β (q β r)) β‘ ((refl β q) β r)
. But both of those
compute to q β r
, so we are done. This computation isnβt
automatic - itβs expressed by the lemma
.
This expresses that the paths behave like morphisms in a category. For a groupoid, we also need inverses and cancellation:
: x β‘ y β y β‘ x
inv {x = x} = J (Ξ» y _ β y β‘ x) refl inv
The operation which assigns inverses has to be involutive, which follows from two computations.
: (p : x β‘ y) β inv (inv p) β‘ p
inv-inv {x = x} =
inv-inv (Ξ» y p β inv (inv p) β‘ p)
J (ap inv (J-refl (Ξ» y _ β y β‘ x) refl) β J-refl (Ξ» y _ β y β‘ x) refl)
And we have to prove that composing with an inverse gives the reflexivity path.
: (p : x β‘ y) β p β inv p β‘ refl
β-invr {x = x} = J (Ξ» y p β p β inv p β‘ refl)
β-invr (β-idl (inv refl) β J-refl (Ξ» y _ β y β‘ x) refl)
: (p : x β‘ y) β inv p β p β‘ refl
β-invl {x = x} = J (Ξ» y p β inv p β p β‘ refl)
β-invl (β-idr (inv refl) β J-refl (Ξ» y _ β y β‘ x) refl)
Cubicallyπ
Now we do the same using hfill
instead of path induction.
module _ where
private variable
: Level
β : Type β
A B : A
w x y z : A
a b c d
open 1Lab.Path
The left and right identity laws follow directly from the two fillers for the composition operation.
: (p : x β‘ y) β p β refl β‘ p
β-idr = sym (β-filler p refl)
β-idr p
: (p : x β‘ y) β refl β p β‘ p
β-idl = sym (β-filler' refl p) β-idl p
For associativity, we use both:
: (p : w β‘ x) (q : x β‘ y) (r : y β‘ z)
β-assoc β p β (q β r) β‘ (p β q) β r
= β-filler p q i β β-filler' q r (~ i) β-assoc p q r i
For cancellation, we need to sketch an open cube where the missing square expresses the equation weβre looking for. Thankfully, we only have to do this once!
: β {x y : A} (p : x β‘ y) β p β sym p β‘ refl
β-invr {x = x} p i j = hcomp (β j β¨ i) Ξ» where
β-invr (k = i0) β p (j β§ ~ i)
k (i = i1) β x
k (j = i0) β x
k (j = i1) β p (~ k β§ ~ i) k
For the other direction, we use the fact that p
is
definitionally equal to sym (sym p)
. In that case, we show
that sym p β sym (sym p) β‘ refl
- which computes to the
thing we want!
: (p : x β‘ y) β sym p β p β‘ refl
β-invl = β-invr (sym p) β-invl p
In addition to the groupoid identities for paths in a type, it has
been established that functions behave like functors: These are the
lemmas ap-refl
, ap-β
and ap-sym
in the 1Lab.Path module.
Convenient helpersπ
Since a lot of Homotopy Type Theory is dealing with paths, this section introduces useful helpers for dealing with -ary compositions. For instance, we know that is , but this involves more than a handful of intermediate steps:
β-cancell: β {β} {A : Type β} {x y z : A} (p : x β‘ y) (q : y β‘ z)
β (sym p β p β q) β‘ q
{y = y} p q i j = hcomp (β i β¨ β j) Ξ» where
β-cancell (k = i0) β p (i β¨ ~ j)
k (i = i0) β β-filler (sym p) (p β q) k j
k (i = i1) β q (j β§ k)
k (j = i0) β y
k (j = i1) β β-fillerβ p q i k
k
β-cancelr: β {β} {A : Type β} {x y z : A} (p : x β‘ y) (q : z β‘ y)
β ((p β sym q) β q) β‘ p
{x = x} {y = y} q p = sym $ β-unique _ Ξ» i j β
β-cancelr (sym p) (~ i) j
β-filler q
commutesβsquare: {p : w β‘ x} {q : w β‘ y} {s : x β‘ z} {r : y β‘ z}
β p β s β‘ q β r
β Square p q s r
{p = p} {q} {s} {r} fill i j =
commutesβsquare (β i β¨ β j) Ξ» where
hcomp (k = i0) β fill j i
k (i = i0) β q (k β§ j)
k (i = i1) β s (~ k β¨ j)
k (j = i0) β β-filler p s (~ k) i
k (j = i1) β β-fillerβ q r k i
k
squareβcommutes: {p : w β‘ x} {q : w β‘ y} {s : x β‘ z} {r : y β‘ z}
β Square p q s r β p β s β‘ q β r
{p = p} {q} {s} {r} fill i j = hcomp (β i β¨ β j) Ξ» where
squareβcommutes (k = i0) β fill j i
k (i = i0) β β-filler p s k j
k (i = i1) β β-fillerβ q r (~ k) j
k (j = i0) β q (~ k β§ i)
k (j = i1) β s (k β¨ i)
k
: {x y z : A} (p : x β‘ y) (q r : y β‘ z)
β-cancel'-l β p β q β‘ p β r β q β‘ r
= sym (β-cancell p q) Β·Β· ap (sym p β_) sq Β·Β· β-cancell p r
β-cancel'-l p q r sq
: {x y z : A} (p : y β‘ z) (q r : x β‘ y)
β-cancel'-r β q β p β‘ r β p β q β‘ r
=
β-cancel'-r p q r sq (β-cancelr q (sym p))
sym (_β sym p) sq
Β·Β· ap (sym p) Β·Β· β-cancelr r
While connections give us
degenerate squares where two adjacent faces are some path and the other
two are refl
, it is sometimes also
useful to have a degenerate square with two pairs of equal adjacent
faces. We can build this using hcomp
and more connections:
double-connection: (p : a β‘ b) (q : b β‘ c)
β Square p p q q
{b = b} p q i j = hcomp (β i β¨ β j) Ξ» where
double-connection (k = i0) β b
k (i = i0) β p (j β¨ ~ k)
k (i = i1) β q (j β§ k)
k (j = i0) β p (i β¨ ~ k)
k (j = i1) β q (i β§ k) k
This corresponds to the following diagram, which expresses the trivial equation :
Groupoid structure of types (cont.)π
A useful fact is that if is a homotopy , then it βcommutes with β, in the following sense:
homotopy-invert: β {a} {A : Type a} {f : A β A}
β (H : (x : A) β f x β‘ x) {x : A}
β H (f x) β‘ ap f (H x)
The proof proceeds entirely using itself, together with the De Morgan algebra structure on the interval.
{f = f} H {x} i j = hcomp (β i β¨ β j) Ξ» where
homotopy-invert (k = i0) β H x (j β§ i)
k (j = i0) β H (f x) (~ k)
k (j = i1) β H x (~ k β§ i)
k (i = i0) β H (f x) (~ k β¨ j)
k (i = i1) β H (H x j) (~ k) k