module Cat.Reasoning {o β} (C : Precategory o β) where
open import Cat.Morphism C public
Reasoning combinators for categoriesπ
When doing category theory, we often have to perform many βtrivialβ algebraic manipulations like reassociation, insertion of identity morphisms, etc. On paper we can omit these steps, but Agda is a bit more picky! We could just do these steps in our proofs one after another, but this clutters things up. Instead, we provide a series of reasoning combinators to make writing (and reading) proofs easier.
Most of these helpers were taken from
agda-categories
.
private variable
: Ob
u v w x y z : Hom x y
a a' a'' b b' b'' c c' c'' d d' d'' e : Hom x y f g g' h h' i
Lensesπ
module _ {w x y z} {a : Hom y z} {b : Hom x y} {c : Hom w x} {f : Hom w z} where abstract
: ((a β b) β c β‘ f) β (a β b β c β‘ f)
reassocl = β-pre-equiv (assoc _ _ _)
reassocl
: (f β‘ (a β b) β c) β (f β‘ a β b β c)
reassocr = β-post-equiv (sym (assoc _ _ _))
reassocr
module reassocl = Equiv reassocl
module reassocr = Equiv reassocr
Identity morphismsπ
abstract
: β {a b} {f : Hom a b} β f β id β‘ id β f
id-comm {f = f} = idr f β sym (idl f)
id-comm
: β {a b} {f : Hom a b} β id β f β‘ f β id
id-comm-sym {f = f} = idl f β sym (idr f)
id-comm-sym
: β {x} β id {x} β id {x} β‘ id
id2 = idl _
id2
: β {a b c} (f : Hom b c) (g : Hom a b) β f β g β id β‘ f β g
idr2 = ap (f β_) (idr g)
idr2 f g
: β {a b c} (f : Hom b c) (g : Hom a b) β (id β f) β g β‘ f β g
idl2 = ap (_β g) (idl f)
idl2 f g
module _ (aβ‘id : a β‘ id) where abstract
: a β f β‘ f
eliml {f = f} =
eliml (_β f) aβ‘id β©
a β f β‘β¨ ap
id β f β‘β¨ idl f β©
f β
: f β a β‘ f
elimr {f = f} =
elimr (f β_) aβ‘id β©
f β a β‘β¨ ap
f β id β‘β¨ idr f β©
f β
: f β a β h β‘ f β h
elim-inner {f = f} = ap (f β_) eliml
elim-inner
: f β‘ a β f
introl = sym eliml
introl
: f β‘ f β a
intror = sym elimr
intror
: f β h β‘ f β a β h
intro-inner {f = f} = ap (f β_) introl intro-inner
Reassocationsπ
We often find ourselves in situations where we have an equality involving the composition of 2 morphisms, but the association is a bit off. These combinators aim to address that situation.
module _ (abβ‘c : a β b β‘ c) where abstract
: a β (b β f) β‘ c β f
pulll {f = f} =
pulll
a β b β f β‘β¨ assoc a b f β©(a β b) β f β‘β¨ ap (_β f) abβ‘c β©
c β f β
: (f β a) β b β‘ f β c
pullr {f = f} =
pullr (f β a) β b β‘Λβ¨ assoc f a b β©
(a β b) β‘β¨ ap (f β_) abβ‘c β©
f β
f β c β
: (f β a) β (b β g) β‘ f β c β g
pull-inner {f = f} = sym (assoc _ _ _) β ap (f β_) pulll
pull-inner
module _ (abcβ‘d : a β b β c β‘ d) where abstract
: a β (b β (c β f)) β‘ d β f
pulll3 {f = f} =
pulll3 (a β_) (assoc _ _ _) β©
a β b β c β f β‘β¨ ap (b β c) β f β‘β¨ assoc _ _ _ β©
a β (a β b β c) β f β‘β¨ ap (_β f) abcβ‘d β©
d β f β
: ((f β a) β b) β c β‘ f β d
pullr3 {f = f} =
pullr3 ((f β a) β b) β c β‘Λβ¨ assoc _ _ _ β©
(f β a) β b β c β‘Λβ¨ assoc _ _ _ β©
(f β_) abcβ‘d β©
f β a β b β c β‘β¨ ap
f β d β
module _ (abcdβ‘e : a β b β c β d β‘ e) where abstract
: a β (b β (c β (d β f))) β‘ e β f
pulll4 {f = f} =
pulll4 (Ξ» x β a β b β x) (assoc _ _ _) β©
a β b β c β d β f β‘β¨ ap (c β d) β f β‘β¨ ap (a β_) (assoc _ _ _) β©
a β b β (b β c β d) β f β‘β¨ assoc _ _ _ β©
a β (a β b β c β d) β f β‘β¨ ap (_β f) abcdβ‘e β©
e β f β
module _ (cβ‘ab : c β‘ a β b) where abstract
: c β f β‘ a β (b β f)
pushl = sym (pulll (sym cβ‘ab))
pushl
: f β c β‘ (f β a) β b
pushr = sym (pullr (sym cβ‘ab))
pushr
: f β c β g β‘ (f β a) β (b β g)
push-inner {f = f} = ap (f β_) pushl β assoc _ _ _
push-inner
module _ (dβ‘abc : d β‘ a β b β c) where abstract
: d β f β‘ a β (b β (c β f))
pushl3 = sym (pulll3 (sym dβ‘abc))
pushl3
: f β d β‘ ((f β a) β b) β c
pushr3 = sym (pullr3 (sym dβ‘abc))
pushr3
module _ (eβ‘abcd : e β‘ a β b β c β d) where abstract
: e β f β‘ a β (b β (c β (d β f)))
pushl4 = sym (pulll4 (sym eβ‘abcd))
pushl4
module _ (p : f β h β‘ g β i) where abstract
: f β (h β b) β‘ g β (i β b)
extendl {b = b} =
extendl (h β b) β‘β¨ assoc f h b β©
f β (f β h) β b β‘β¨ ap (_β b) p β©
(g β i) β b β‘Λβ¨ assoc g i b β©
(i β b) β
g β
: (a β f) β h β‘ (a β g) β i
extendr {a = a} =
extendr (a β f) β h β‘Λβ¨ assoc a f h β©
(f β h) β‘β¨ ap (a β_) p β©
a β (g β i) β‘β¨ assoc a g i β©
a β (a β g) β i β
: a β f β h β b β‘ a β g β i β b
extend-inner {a = a} = ap (a β_) extendl
extend-inner
module _ (p : a β b β c β‘ d β f β g) where abstract
: a β (b β (c β h)) β‘ d β (f β (g β h))
extendl3 = pulll3 p β sym (pulll3 refl)
extendl3
: ((h β a) β b) β c β‘ ((h β d) β f) β g
extendr3 = pullr3 p β sym (pullr3 refl)
extendr3
module _ (p : a β b β c β d β‘ e β f β g β h) where abstract
: a β b β c β d β i β‘ e β f β g β h β i
extendl4 = pulll4 p β sym (pulll4 refl) extendl4
We also define some useful combinators for performing repeated pulls/pushes.
abstract
centralize: f β g β‘ a β b
β h β i β‘ c β d
β f β g β h β i β‘ a β (b β c) β d
{f = f} {g = g} {a = a} {b = b} {h = h} {i = i} {c = c} {d = d} p q =
centralize
f β g β h β i β‘β¨ pulll p β©(a β b) β h β i β‘β¨ pullr (pushr q) β©
(b β c) β d β
a β
centralizel: f β g β‘ a β b
β f β g β h β i β‘ a β (b β h) β i
= centralize p refl
centralizel p
centralizer: h β i β‘ c β d
β f β g β h β i β‘ f β (g β c) β d
= centralize refl p
centralizer p
disperse: f β g β‘ a β b
β h β i β‘ c β d
β f β (g β h) β i β‘ a β b β c β d
{f = f} {g = g} {a = a} {b = b} {h = h} {i = i} {c = c} {d = d} p q =
disperse (g β h) β i β‘β¨ pushr (pullr q) β©
f β (f β g) β c β d β‘β¨ pushl p β©
a β b β c β d β
dispersel: f β g β‘ a β b
β f β (g β h) β i β‘ a β b β h β i
= disperse p refl
dispersel p
disperser: h β i β‘ c β d
β f β (g β h) β i β‘ f β g β c β d
= disperse refl p disperser p
Cancellationπ
These lemmas do 2 things at once: rearrange parenthesis, and also
remove things that are equal to id
.
module _ (inv : h β i β‘ id) where abstract
: h β (i β f) β‘ f
cancell {f = f} =
cancell (i β f) β‘β¨ pulll inv β©
h β
id β f β‘β¨ idl f β©
f β
: (f β h) β i β‘ f
cancelr {f = f} =
cancelr (f β h) β i β‘β¨ pullr inv β©
f β id β‘β¨ idr f β©
f β
: f β‘ h β (i β f)
insertl = sym cancell
insertl
: f β‘ (f β h) β i
insertr = sym cancelr
insertr
: (f β h) β (i β g) β‘ f β g
cancel-inner = pulll cancelr
cancel-inner
: f β g β‘ (f β h) β (i β g)
insert-inner = pushl insertr
insert-inner
: (f β g β h) β i β‘ f β g
deleter = pullr cancelr
deleter
: h β (i β f) β g β‘ f β g
deletel = pulll cancell
deletel
module _ (inv : g β h β i β‘ id) where abstract
: g β (h β (i β f)) β‘ f
cancell3 {f = f} = pulll3 inv β idl f
cancell3
: ((f β g) β h) β i β‘ f
cancelr3 {f = f} = pullr3 inv β idr f
cancelr3
: f β‘ g β (h β (i β f))
insertl3 = sym cancell3
insertl3
: f β‘ ((f β g) β h) β i
insertr3 = sym cancelr3 insertr3
We also have combinators which combine expanding on one side with a cancellation on the other side:
: g β‘ h β i β f β h β‘ id β f β g β‘ i
lswizzle {g = g} {h = h} {i = i} {f = f} p q =
lswizzle _β_ refl p β©
f β g β‘β¨ apβ
f β h β i β‘β¨ cancell q β©
i β
: g β‘ i β h β h β f β‘ id β g β f β‘ i
rswizzle {g = g} {i = i} {h = h} {f = f} p q =
rswizzle _β_ p refl β©
g β f β‘β¨ apβ (i β h) β f β‘β¨ cancelr q β©
i β
The following βswizzleβ operation can be pictured as flipping a commutative square along an axis, provided the morphisms on that axis are invertible.
: f β g β‘ h β i β g β g' β‘ id β h' β h β‘ id β h' β f β‘ i β g'
swizzle {f = f} {g = g} {h = h} {i = i} {g' = g'} {h' = h'} p q r =
swizzle (sym (assoc _ _ _ β rswizzle (sym p) q)) r lswizzle
Isomorphismsπ
These lemmas are useful for proving that partial inverses to an isomorphism are unique. Thereβs a helper for proving uniqueness of left inverses, of right inverses, and for proving that any left inverse must match any right inverse.
module _ {y z} (f : y β
z) where abstract
open _β
_
left-inv-unique: β {g h}
β g β f .to β‘ id β h β f .to β‘ id
β g β‘ h
{g = g} {h = h} p q =
left-inv-unique (f .invl) β©
g β‘β¨ intror .to β f .from β‘β¨ extendl (p β sym q) β©
g β f .to β f .from β‘β¨ elimr (f .invl) β©
h β f
h β
left-right-inv-unique: β {g h}
β g β f .to β‘ id β f .to β h β‘ id
β g β‘ h
{g = g} {h = h} p q =
left-right-inv-unique (f .invl) β©
g β‘β¨ intror .to β f .from β‘β¨ cancell p β©
g β f .from β‘β¨ intror q β©
f .from β f .to β h β‘β¨ cancell (f .invr) β©
f
h β
right-inv-unique: β {g h}
β f .to β g β‘ id β f .to β h β‘ id
β g β‘ h
{g = g} {h} p q =
right-inv-unique (f .invr) β©
g β‘β¨ introl (f .from β f .to) β g β‘β¨ pullr (p β sym q) β©
.from β f .to β h β‘β¨ cancell (f .invr) β©
f h β
Lenses for isomorphismsπ
module _
{x y z} {a : Hom x z} {f : Hom x y} {b : Hom y z}
(f-inv : is-invertible f)
where abstract
private module f = is-invertible f-inv
: (a β f.inv β‘ b) β (a β‘ b β f)
pre-invr =
pre-invr (ap (_β f) , equivβcancellable (invertible-precomp-equiv f-inv))
(insertr f.invr)
βe β-pre-equiv
: (b β‘ a β f.inv) β (b β f β‘ a)
post-invr = sym-equiv βe pre-invr βe sym-equiv
post-invr
module pre-invr = Equiv pre-invr
module post-invr = Equiv post-invr
module _
{w x y} {a : Hom w y} {f : Hom x y} {b : Hom w x}
(f-inv : is-invertible f)
where abstract
private module f = is-invertible f-inv
: (f.inv β a β‘ b) β (a β‘ f β b)
pre-invl =
pre-invl (ap (f β_) , equivβcancellable (invertible-postcomp-equiv f-inv))
(insertl f.invl)
βe β-pre-equiv
: (b β‘ f.inv β a) β (f β b β‘ a)
post-invl = sym-equiv βe pre-invl βe sym-equiv
post-invl
module pre-invl = Equiv pre-invl
module post-invl = Equiv post-invl
If we have a commuting triangle of isomorphisms, then we can flip one of the sides to obtain a new commuting triangle of isomorphisms.
:
Iso-swapr β {a : x β
y} {b : y β
z} {c : x β
z}
β b βIso a β‘ c
β a β‘ b Isoβ»ΒΉ βIso c
{a = a} {b = b} {c = c} p = β
-path $
Iso-swapr .to β‘β¨ introl (b .invr) β©
a (b .from β b .to) β a .to β‘β¨ pullr (ap to p) β©
.from β c .to β
b
:
Iso-swapl β {a : x β
y} {b : y β
z} {c : x β
z}
β b βIso a β‘ c
β b β‘ c βIso a Isoβ»ΒΉ
{a = a} {b = b} {c = c} p = β
-path $
Iso-swapl .to β‘β¨ intror (a .invl) β©
b .to β a .to β a .from β‘β¨ pulll (ap to p) β©
b .to β a .from β c
Assume we have a prism of isomorphisms, as in the following diagram:
If the top, front, left, and right faces all commute, then so does the bottom face.
: β {a : u β
v} {b : v β
w} {c : u β
w}
Iso-prism β {d : u β
x} {e : v β
y} {f : w β
z}
β {g : x β
y} {h : y β
z} {i : x β
z}
β b βIso a β‘ c
β e βIso a β‘ g βIso d
β f βIso b β‘ h βIso e
β f βIso c β‘ i βIso d
β h βIso g β‘ i
{a = a} {b} {c} {d} {e} {f} {g} {h} {i} top left right front =
Iso-prism
β
-path $.to β g .to β‘β¨ apβ _β_ (ap to (Iso-swapl (sym right))) (ap to (Iso-swapl (sym left)) β sym (assoc _ _ _)) β©
h ((f .to β b .to) β e .from) β (e .to β a .to β d .from) β‘β¨ cancel-inner (e .invr) β©
(f .to β b .to) β (a .to β d .from) β‘β¨ pull-inner (ap to top) β©
.to β c .to β d .from β‘β¨ assoc _ _ _ β sym (ap to (Iso-swapl (sym front))) β©
f .to β i
Notationπ
When doing equational reasoning, itβs often somewhat clumsy to have
to write ap (f β_) p
when proving that
f β g β‘ f β h
. To fix this, we steal some cute mixfix
notation from agda-categories
which allows us to write
β‘β¨ reflβ©ββ¨ p β©
instead, which is much more aesthetically
pleasing!
_β©ββ¨_ : f β‘ h β g β‘ i β f β g β‘ h β i
_β©ββ¨_ = apβ _β_
infixr 20 _β©ββ¨_
_ : g β‘ h β f β g β‘ f β h
reflβ©ββ¨_ {f = f} p = ap (f β_) p
reflβ©ββ¨
_β©ββ¨refl : f β‘ h β f β g β‘ h β g
_β©ββ¨refl {g = g} p = ap (_β g) p
infix 21 reflβ©ββ¨_
infix 22 _β©ββ¨refl