module Cat.Reasoning {o β} (C : Precategory o β) where
open import Cat.Morphism C publicReasoning 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
  u v w x y z : Ob
  a a' a'' b b' b'' c c' c'' d d' d'' e : Hom x y
  f g g' h h' i : Hom x yLensesπ
module _ {w x y z} {a : Hom y z} {b : Hom x y} {c : Hom w x} {f : Hom w z} where abstract
  reassocl : ((a β b) β c β‘ f) β (a β b β c β‘ f)
  reassocl = β-pre-equiv (assoc _ _ _)
  reassocr : (f β‘ (a β b) β c) β (f β‘ a β b β c)
  reassocr = β-post-equiv (sym (assoc _ _ _))
  module reassocl = Equiv reassocl
  module reassocr = Equiv reassocrIdentity morphismsπ
abstract
  id-comm : β {a b} {f : Hom a b} β f β id β‘ id β f
  id-comm {f = f} = idr f β sym (idl f)
  id-comm-sym : β {a b} {f : Hom a b} β id β f β‘ f β id
  id-comm-sym {f = f} = idl f β sym (idr f)
  id2 : β {x} β id {x} β id {x} β‘ id
  id2 = idl _
  idr2 : β {a b c} (f : Hom b c) (g : Hom a b) β f β g β id β‘ f β g
  idr2 f g = ap (f β_) (idr g)
  idl2 : β {a b c} (f : Hom b c) (g : Hom a b) β (id β f) β g β‘ f β g
  idl2 f g = ap (_β g) (idl f)
module _ (aβ‘id : a β‘ id) where abstract
  eliml : a β f β‘ f
  eliml {f = f} =
    a β f β‘β¨ ap (_β f) aβ‘id β©
    id β f β‘β¨ idl f β©
    f β
  elimr : f β a β‘ f
  elimr {f = f} =
    f β a β‘β¨ ap (f β_) aβ‘id β©
    f β id β‘β¨ idr f β©
    f β
  elim-inner : f β a β h β‘ f β h
  elim-inner {f = f} = ap (f β_) eliml
  introl : f β‘ a β f
  introl = sym eliml
  intror : f β‘ f β a
  intror = sym elimr
  intro-inner : f β h β‘ f β a β h
  intro-inner {f = f} = ap (f β_) introlReassocationsπ
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
  pulll : a β (b β f) β‘ c β f
  pulll {f = f} =
    a β b β f   β‘β¨ assoc a b f β©
    (a β b) β f β‘β¨ ap (_β f) abβ‘c β©
    c β f β
  pullr : (f β a) β b β‘ f β c
  pullr {f = f} =
    (f β a) β b β‘Λβ¨ assoc f a b β©
    f β (a β b) β‘β¨ ap (f β_) abβ‘c β©
    f β c β
  pull-inner : (f β a) β (b β g) β‘ f β c β g
  pull-inner {f = f} = sym (assoc _ _ _) β ap (f β_) pulll
module _ (abcβ‘d : a β b β c β‘ d) where abstract
  pulll3 : a β (b β (c β f)) β‘ d β f
  pulll3 {f = f} =
    a β b β c β f   β‘β¨ ap (a β_) (assoc _ _ _) β©
    a β (b β c) β f β‘β¨ assoc _ _ _ β©
    (a β b β c) β f β‘β¨ ap (_β f) abcβ‘d β©
    d β f           β
  pullr3 : ((f β a) β b) β c β‘ f β d
  pullr3 {f = f} =
    ((f β a) β b) β c β‘Λβ¨ assoc _ _ _ β©
    (f β a) β b β c   β‘Λβ¨ assoc _ _ _ β©
    f β a β b β c     β‘β¨ ap (f β_) abcβ‘d β©
    f β d β
module _ (abcdβ‘e : a β b β c β d β‘ e) where abstract
  pulll4 : a β (b β (c β (d β f))) β‘ e β f
  pulll4 {f = f} =
    a β b β c β d β f β‘β¨ ap (Ξ» x β a β b β x) (assoc _ _ _) β©
    a β b β (c β d) β f β‘β¨ ap (a β_) (assoc _ _ _) β©
    a β (b β c β d) β f β‘β¨ assoc _ _ _ β©
    (a β b β c β d) β f β‘β¨ ap (_β f) abcdβ‘e β©
    e β f β
module _ (cβ‘ab : c β‘ a β b) where abstract
  pushl : c β f β‘ a β (b β f)
  pushl = sym (pulll (sym cβ‘ab))
  pushr : f β c β‘ (f β a) β b
  pushr = sym (pullr (sym cβ‘ab))
  push-inner : f β c β g β‘ (f β a) β (b β g)
  push-inner {f = f} = ap (f β_) pushl β assoc _ _ _
module _ (dβ‘abc : d β‘ a β b β c) where abstract
  pushl3 : d β f β‘ a β (b β (c β f))
  pushl3 = sym (pulll3 (sym dβ‘abc))
  pushr3 : f β d β‘ ((f β a) β b) β c
  pushr3 = sym (pullr3 (sym dβ‘abc))
module _ (eβ‘abcd : e β‘ a β b β c β d) where abstract
  pushl4 : e β f β‘ a β (b β (c β (d β f)))
  pushl4 = sym (pulll4 (sym eβ‘abcd))
module _ (p : f β h β‘ g β i) where abstract
  extendl : f β (h β b) β‘ g β (i β b)
  extendl {b = b} =
    f β (h β b) β‘β¨ assoc f h b β©
    (f β h) β b β‘β¨ ap (_β b) p β©
    (g β i) β b β‘Λβ¨ assoc g i b β©
    g β (i β b) β
  extendr : (a β f) β h β‘ (a β g) β i
  extendr {a = a} =
    (a β f) β h β‘Λβ¨ assoc a f h β©
    a β (f β h) β‘β¨ ap (a β_) p β©
    a β (g β i) β‘β¨ assoc a g i β©
    (a β g) β i β
  extend-inner : a β f β h β b β‘ a β g β i β b
  extend-inner {a = a} = ap (a β_) extendl
module _ (p : a β b β c β‘ d β f β g) where abstract
  extendl3 : a β (b β (c β h)) β‘ d β (f β (g β h))
  extendl3 = pulll3 p β sym (pulll3 refl)
  extendr3 : ((h β a) β b) β c β‘ ((h β d) β f) β g
  extendr3 = pullr3 p β sym (pullr3 refl)
module _ (p : a β b β c β d β‘ e β f β g β h) where abstract
  extendl4 : a β b β c β d β i β‘ e β f β g β h β i
  extendl4 = pulll4 p β sym (pulll4 refl)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
  centralize {f = f} {g = g} {a = a} {b = b} {h = h} {i = i} {c = c} {d = d} p q =
    f β g β h β i   β‘β¨ pulll p β©
    (a β b) β h β i β‘β¨ pullr (pushr q) β©
    a β (b β c) β d β
  centralizel
    : f β g β‘ a β b
    β f β g β h β i β‘ a β (b β h) β i
  centralizel p = centralize p refl
  centralizer
    : h β i β‘ c β d
    β f β g β h β i β‘ f β (g β c) β d
  centralizer p = centralize refl p
  disperse
    : f β g β‘ a β b
    β h β i β‘ c β d
    β f β (g β h) β i β‘ a β b β c β d
  disperse {f = f} {g = g} {a = a} {b = b} {h = h} {i = i} {c = c} {d = d} p q =
    f β (g β h) β i β‘β¨ pushr (pullr q) β©
    (f β g) β c β d β‘β¨ pushl p β©
    a β b β c β d β
  dispersel
    : f β g β‘ a β b
    β f β (g β h) β i β‘ a β b β h β i
  dispersel p = disperse p refl
  disperser
    : h β i β‘ c β d
    β f β (g β h) β i β‘ f β g β c β d
  disperser p = disperse refl pCancellationπ
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
  cancell : h β (i β f) β‘ f
  cancell {f = f} =
    h β (i β f) β‘β¨ pulll inv β©
    id β f      β‘β¨ idl f β©
    f           β
  cancelr : (f β h) β i β‘ f
  cancelr {f = f} =
    (f β h) β i β‘β¨ pullr inv β©
    f β id      β‘β¨ idr f β©
    f           β
  insertl : f β‘ h β (i β f)
  insertl = sym cancell
  insertr : f β‘ (f β h) β i
  insertr = sym cancelr
  cancel-inner : (f β h) β (i β g) β‘ f β g
  cancel-inner = pulll cancelr
  insert-inner : f β g β‘ (f β h) β (i β g)
  insert-inner = pushl insertr
  deleter : (f β g β h) β i β‘ f β g
  deleter = pullr cancelr
  deletel : h β (i β f) β g β‘ f β g
  deletel = pulll cancell
module _ (inv : g β h β i β‘ id) where abstract
  cancell3 : g β (h β (i β f)) β‘ f
  cancell3 {f = f} = pulll3 inv β idl f
  cancelr3 : ((f β g) β h) β i β‘ f
  cancelr3 {f = f} = pullr3 inv β idr f
  insertl3 : f β‘ g β (h β (i β f))
  insertl3 = sym cancell3
  insertr3 : f β‘ ((f β g) β h) β i
  insertr3 = sym cancelr3We also have combinators which combine expanding on one side with a cancellation on the other side:
lswizzle : g β‘ h β i β f β h β‘ id β f β g β‘ i
lswizzle {g = g} {h = h} {i = i} {f = f} p q =
  f β g     β‘β¨ apβ _β_ refl p β©
  f β h β i β‘β¨ cancell q β©
  i         β
rswizzle : g β‘ i β h β h β f β‘ id β g β f β‘ i
rswizzle {g = g} {i = i} {h = h} {f = f} p q =
  g β f       β‘β¨ apβ _β_ p refl β©
  (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.
swizzle : 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 =
  lswizzle (sym (assoc _ _ _ β rswizzle (sym p) q)) rSections and Retractsπ
We can repackage our βswizzlingβ lemmas to move around sections and retracts.
module _
  {f : Hom x y}
  (f-section : has-section f)
  where abstract
  private module f = has-section f-section
  pre-section : a β f β‘ b β a β‘ b β f.section
  pre-section {a = a} {b = b} p = sym (rswizzle (sym p) f.is-section)
  post-section : f.section β a β‘ b β a β‘ f β b
  post-section {a = a} {b = b} p = sym (lswizzle (sym p) f.is-section)
module _
  {f : Hom x y}
  (f-retract : has-retract f)
  where abstract
  private module f = has-retract f-retract
  pre-retract : a β f.retract β‘ b β a β‘ b β f
  pre-retract {a = a} {b = b} p = sym (rswizzle (sym p) f.is-retract)
  post-retract : f β a β‘ b β a β‘ f.retract β b
  post-retract {a = a} {b = b} p = sym (lswizzle (sym p) f.is-retract)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
  open _β
_
  abstract
    left-inv-unique
      : β {g h}
      β g β f .to β‘ id β h β f .to β‘ id
      β g β‘ h
    left-inv-unique {g = g} {h = h} p q =
      g                   β‘β¨ intror (f .invl) β©
      g β f .to β f .from β‘β¨ extendl (p β sym q) β©
      h β f .to β f .from β‘β¨ elimr (f .invl) β©
      h                   β
    left-right-inv-unique
      : β {g h}
      β g β f .to β‘ id β f .to β h β‘ id
      β g β‘ h
    left-right-inv-unique {g = g} {h = h} p q =
      g                    β‘β¨ intror (f .invl) β©
      g β f .to β f .from  β‘β¨ cancell p β©
      f .from              β‘β¨ intror q β©
      f .from β f .to β h  β‘β¨ cancell (f .invr) β©
      h                    β
    right-inv-unique
      : β {g h}
      β f .to β g β‘ id β f .to β h β‘ id
      β g β‘ h
    right-inv-unique {g = g} {h} p q =
      g                     β‘β¨ introl (f .invr) β©
      (f .from β f .to) β g β‘β¨ pullr (p β sym q) β©
      f .from β f .to β h   β‘β¨ cancell (f .invr) β©
      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
  pre-invr : (a β f.inv β‘ b) β (a β‘ b β f)
  pre-invr =
    (ap (_β f) , equivβcancellable (invertible-precomp-equiv f-inv))
    βe β-pre-equiv (insertr f.invr)
  post-invr : (b β‘ a β f.inv) β (b β f β‘ a)
  post-invr = sym-equiv βe pre-invr βe sym-equiv
  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
  pre-invl : (f.inv β a β‘ b) β (a β‘ f β b)
  pre-invl =
    (ap (f β_) , equivβcancellable (invertible-postcomp-equiv f-inv))
    βe β-pre-equiv (insertl f.invl)
  post-invl : (b β‘ f.inv β a) β (f β b β‘ a)
  post-invl = sym-equiv βe pre-invl βe sym-equiv
  module pre-invl = Equiv pre-invl
  module post-invl = Equiv post-invlIf 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
Iso-swapr {a = a} {b = b} {c = c} p = β
-path $
  a .to                     β‘β¨ introl (b .invr) β©
  (b .from β b .to) β a .to β‘β¨ pullr (ap to p) β©
  b .from β c .to           β
Iso-swapl :
  β {a : x β
 y} {b : y β
 z} {c : x β
 z}
  β b βIso a β‘ c
  β b β‘ c βIso a Isoβ»ΒΉ
Iso-swapl {a = a} {b = b} {c = c} p = β
-path $
  b .to                   β‘β¨ intror (a .invl) β©
  b .to β a .to β a .from β‘β¨ pulll (ap to p) β©
  c .to β a .from         β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.
Iso-prism : β {a : u β
 v} {b : v β
 w} {c : u β
 w}
      β {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
Iso-prism {a = a} {b} {c} {d} {e} {f} {g} {h} {i} top left right front =
  β
-path $
    h .to β g .to                                           β‘β¨ apβ _β_ (ap to (Iso-swapl (sym right))) (ap to (Iso-swapl (sym left)) β sym (assoc _ _ _)) β©
    ((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) β©
    f .to β c .to β d .from                                 β‘β¨ assoc _ _ _ β sym (ap to (Iso-swapl (sym front))) β©
    i .to β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 _β©ββ¨_
reflβ©ββ¨_ : g β‘ h β f β g β‘ f β h
reflβ©ββ¨_ {f = f} p = ap (f β_) p
_β©ββ¨refl : f β‘ h β f β g β‘ h β g
_β©ββ¨refl {g = g} p = ap (_β g) p
infix 21 reflβ©ββ¨_
infix 22 _β©ββ¨refl