Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

Right adjoint to Functor.Slice.Free #408

Draft
wants to merge 8 commits into
base: master
Choose a base branch
from
Draft
Show file tree
Hide file tree
Changes from 6 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
4 changes: 2 additions & 2 deletions src/Categories/Adjoint/Instance/BaseChange.agda
Original file line number Diff line number Diff line change
Expand Up @@ -6,7 +6,7 @@ module Categories.Adjoint.Instance.BaseChange {o ℓ e} (C : Category o ℓ e) w

open import Categories.Adjoint using (_⊣_)
open import Categories.Adjoint.Compose using (_∘⊣_)
open import Categories.Adjoint.Instance.Slice using (Forgetful⊣Free)
open import Categories.Adjoint.Instance.Slice using (TotalSpace⊣ConstantFamily)
open import Categories.Category.Slice C using (Slice)
open import Categories.Category.Slice.Properties C using (pullback⇒product; slice-slice≃slice)
open import Categories.Category.Equivalence.Properties using (module C≅D)
Expand All @@ -18,4 +18,4 @@ open Category C
module _ {A B : Obj} (f : B ⇒ A) (pullback : ∀ {C} {h : C ⇒ A} → Pullback f h) where

!⊣* : BaseChange! f ⊣ BaseChange* f pullback
!⊣* = C≅D.L⊣R (slice-slice≃slice f) ∘⊣ Forgetful⊣Free (Slice A) (pullback⇒product pullback)
!⊣* = C≅D.L⊣R (slice-slice≃slice f) ∘⊣ TotalSpace⊣ConstantFamily (Slice A) (pullback⇒product pullback)
19 changes: 9 additions & 10 deletions src/Categories/Adjoint/Instance/Slice.agda
Original file line number Diff line number Diff line change
Expand Up @@ -9,7 +9,7 @@ open import Categories.Category.BinaryProducts C
open import Categories.Category.Cartesian using (Cartesian)
open import Categories.Category.CartesianClosed using (CartesianClosed)
open import Categories.Category.Slice C using (SliceObj; sliceobj; Slice⇒; slicearr)
open import Categories.Functor.Slice C using (Forgetful; Free; Coforgetful)
open import Categories.Functor.Slice C using (TotalSpace; ConstantFamily; InternalSections)
open import Categories.Morphism.Reasoning C
open import Categories.NaturalTransformation using (ntHelper)
open import Categories.Diagram.Pullback C hiding (swap)
Expand All @@ -30,8 +30,8 @@ module _ {A : Obj} (product : ∀ {X} → Product A X) where
module product {X} = Product (product {X})
open product

Forgetful⊣Free : ForgetfulFree product
Forgetful⊣Free = record
TotalSpace⊣ConstantFamily : TotalSpaceConstantFamily product
TotalSpace⊣ConstantFamily = record
{ unit = ntHelper record
{ η = λ _ → slicearr project₁
; commute = λ {X} {Y} f → begin
Expand Down Expand Up @@ -60,16 +60,16 @@ module _ {A : Obj} (ccc : CartesianClosed C) (pullback : ∀ {X} {Y} {Z} (h : X
open Terminal terminal
open BinaryProducts products

Free⊣Coforgetful : Free {A = A} product ⊣ Coforgetful ccc pullback
Free⊣Coforgetful = record
ConstantFamily⊣InternalSections : ConstantFamily {A = A} product ⊣ InternalSections ccc pullback
ConstantFamily⊣InternalSections = record
{ unit = ntHelper record
{ η = λ X → p.universal (sliceobj π₁) (λ-unique₂′ (unit-pb X))
; commute = λ {S} {T} f → p.unique-diagram (sliceobj π₁) !-unique₂ $ begin
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

For these longer proofs, best to put them either in where or named in a private block. Purely for efficiency.

p.p₂ (sliceobj π₁) ∘ p.universal (sliceobj π₁) _ ∘ f ≈⟨ pullˡ (p.p₂∘universal≈h₂ (sliceobj π₁)) ⟩
λg swap ∘ f ≈⟨ subst ⟩
λg (swap ∘ first f) ≈⟨ λ-cong swap∘⁂ ⟩
λg (second f ∘ swap) ≈˘⟨ λ-cong (∘-resp-≈ʳ β′) ⟩
λg (second f ∘ eval′ ∘ first (λg swap)) ≈˘⟨ λ-cong (∘-resp-≈ʳ (∘-resp-≈ʳ (-cong (p.p₂∘universal≈h₂ (sliceobj π₁)) Equiv.refl))) ⟩
λg (second f ∘ eval′ ∘ first (λg swap)) ≈˘⟨ λ-cong (∘-resp-≈ʳ (∘-resp-≈ʳ (first-cong (p.p₂∘universal≈h₂ (sliceobj π₁))))) ⟩
λg (second f ∘ eval′ ∘ first (p.p₂ (sliceobj π₁) ∘ p.universal (sliceobj π₁) _)) ≈˘⟨ λ-cong (pull-last first∘first) ⟩
λg ((second f ∘ eval′ ∘ first (p.p₂ (sliceobj π₁))) ∘ first (p.universal (sliceobj π₁) _)) ≈˘⟨ subst ⟩
λg (second f ∘ eval′ ∘ first (p.p₂ (sliceobj π₁))) ∘ p.universal (sliceobj π₁) _ ≈˘⟨ pullˡ (p.p₂∘universal≈h₂ (sliceobj π₁)) ⟩
Expand All @@ -80,7 +80,7 @@ module _ {A : Obj} (ccc : CartesianClosed C) (pullback : ∀ {X} {Y} {Z} (h : X
; commute = λ {S} {T} f → begin
(eval′ ∘ first (p.p₂ T) ∘ swap) ∘ second (p.universal T _) ≈⟨ pull-last swap∘⁂ ⟩
eval′ ∘ first (p.p₂ T) ∘ first (p.universal T _) ∘ swap ≈⟨ refl⟩∘⟨ pullˡ first∘first ⟩
eval′ ∘ first (p.p₂ T ∘ p.universal T _) ∘ swap ≈⟨ refl⟩∘⟨ -cong (p.p₂∘universal≈h₂ T) Equiv.refl ⟩∘⟨refl ⟩
eval′ ∘ first (p.p₂ T ∘ p.universal T _) ∘ swap ≈⟨ refl⟩∘⟨ first-cong (p.p₂∘universal≈h₂ T) ⟩∘⟨refl ⟩
eval′ ∘ first (λg (h f ∘ eval′ ∘ first (p.p₂ S))) ∘ swap ≈⟨ pullˡ β′ ⟩
(h f ∘ eval′ ∘ first (p.p₂ S)) ∘ swap ≈⟨ assoc²' ⟩
h f ∘ eval′ ∘ first (p.p₂ S) ∘ swap ∎
Expand All @@ -89,7 +89,7 @@ module _ {A : Obj} (ccc : CartesianClosed C) (pullback : ∀ {X} {Y} {Z} (h : X
(eval′ ∘ first (p.p₂ (sliceobj π₁)) ∘ swap) ∘ second (p.universal (sliceobj π₁) _) ≈⟨ assoc²' ⟩
eval′ ∘ first (p.p₂ (sliceobj π₁)) ∘ swap ∘ second (p.universal (sliceobj π₁) _) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ swap∘⁂ ⟩
eval′ ∘ first (p.p₂ (sliceobj π₁)) ∘ first (p.universal (sliceobj π₁) _) ∘ swap ≈⟨ refl⟩∘⟨ pullˡ first∘first ⟩
eval′ ∘ first (p.p₂ (sliceobj π₁) ∘ p.universal (sliceobj π₁) _) ∘ swap ≈⟨ refl⟩∘⟨ -cong (p.p₂∘universal≈h₂ (sliceobj π₁)) Equiv.refl ⟩∘⟨refl ⟩
eval′ ∘ first (p.p₂ (sliceobj π₁) ∘ p.universal (sliceobj π₁) _) ∘ swap ≈⟨ refl⟩∘⟨ first-cong (p.p₂∘universal≈h₂ (sliceobj π₁)) ⟩∘⟨refl ⟩
eval′ ∘ first (λg swap) ∘ swap ≈⟨ pullˡ β′ ⟩
swap ∘ swap ≈⟨ swap∘swap ⟩
id ∎
Expand All @@ -99,8 +99,7 @@ module _ {A : Obj} (ccc : CartesianClosed C) (pullback : ∀ {X} {Y} {Z} (h : X
λg ((eval′ ∘ first (p.p₂ X) ∘ swap) ∘ eval′) ∘ p.p₂ (sliceobj π₁) ∘ p.universal (sliceobj π₁) _ ≈⟨ refl⟩∘⟨ p.p₂∘universal≈h₂ (sliceobj π₁) ⟩
λg ((eval′ ∘ first (p.p₂ X) ∘ swap) ∘ eval′) ∘ λg swap ≈⟨ subst ⟩
λg (((eval′ ∘ first (p.p₂ X) ∘ swap) ∘ eval′) ∘ first (λg swap)) ≈⟨ λ-cong (pullʳ β′) ⟩
λg ((eval′ ∘ first (p.p₂ X) ∘ swap) ∘ swap) ≈⟨ λ-cong (pull-last swap∘swap) ⟩
λg (eval′ ∘ first (p.p₂ X) ∘ id) ≈⟨ λ-cong (∘-resp-≈ʳ identityʳ) ⟩
λg ((eval′ ∘ first (p.p₂ X) ∘ swap) ∘ swap) ≈⟨ λ-cong (pullʳ (cancelʳ swap∘swap)) ⟩
λg (eval′ ∘ first (p.p₂ X)) ≈⟨ η-exp′ ⟩
p.p₂ X ≈˘⟨ identityʳ ⟩
p.p₂ X ∘ id ∎
Expand Down
6 changes: 6 additions & 0 deletions src/Categories/Category/BinaryProducts.agda
Original file line number Diff line number Diff line change
Expand Up @@ -94,6 +94,12 @@ record BinaryProducts : Set (levelOfTerm 𝒞) where
⁂-cong₂ : f ≈ g → h ≈ i → f ⁂ h ≈ g ⁂ i
⁂-cong₂ = [ product ⇒ product ]×-cong₂

first-cong : f ≈ g → f ⁂ h ≈ g ⁂ h
JacquesCarette marked this conversation as resolved.
Show resolved Hide resolved
first-cong f≈g = ⁂-cong₂ f≈g Equiv.refl

second-cong : g ≈ h → f ⁂ g ≈ f ⁂ h
second-cong g≈h = ⁂-cong₂ Equiv.refl g≈h

⁂∘⟨⟩ : (f ⁂ g) ∘ ⟨ f′ , g′ ⟩ ≈ ⟨ f ∘ f′ , g ∘ g′ ⟩
⁂∘⟨⟩ = [ product ⇒ product ]×∘⟨⟩

Expand Down
4 changes: 2 additions & 2 deletions src/Categories/Category/CartesianClosed.agda
Original file line number Diff line number Diff line change
Expand Up @@ -56,7 +56,7 @@ record CartesianClosed : Set (levelOfTerm 𝒞) where
open CartesianMonoidal cartesian using (A×⊤≅A)
open BinaryProducts cartesian.products using (_×_; product; π₁; π₂; ⟨_,_⟩;
project₁; project₂; η; ⟨⟩-cong₂; ⟨⟩∘; _⁂_; ⟨⟩-congˡ; ⟨⟩-congʳ;
first∘first; firstid; first; second; first↔second; second∘second; -cong; -×_)
first∘first; firstid; first; second; first↔second; second∘second; second-cong; -×_)
open Terminal cartesian.terminal using (⊤; !; !-unique₂; ⊤-id)

B^A×A : ∀ B A → Product (B ^ A) A
Expand Down Expand Up @@ -194,7 +194,7 @@ record CartesianClosed : Set (levelOfTerm 𝒞) where
; identity = λ-cong (identityˡ ○ (elimʳ (id×id product))) ○ η-id′
; homomorphism = λ-unique₂′ helper
; F-resp-≈ = λ where
(eq₁ , eq₂) → λ-cong (∘-resp-≈ eq₂ (∘-resp-≈ʳ (-cong₂ refl eq₁)))
(eq₁ , eq₂) → λ-cong (∘-resp-≈ eq₂ (∘-resp-≈ʳ (second-cong eq₁)))
}
where helper : eval′ ∘ first (λg ((g ∘ f) ∘ eval′ ∘ second (h ∘ i)))
≈ eval′ ∘ first (λg (g ∘ eval′ ∘ second i) ∘ λg (f ∘ eval′ ∘ second h))
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -55,7 +55,7 @@ module FromCartesian o′ ℓ′ {o ℓ e} {C : Category o ℓ e} (Car : Cartesi
in begin
F.₁ (second (f C.∘ g)) ⟨$⟩ x ≈˘⟨ [ F ]-resp-∘ second∘second ⟩
F.₁ (second g) ⟨$⟩ (F.₁ (second f) ⟨$⟩ x) ∎
; F-resp-≈ = λ {Y Z} {f g} eq → F.F-resp-≈ (-cong₂ C.Equiv.refl eq)
; F-resp-≈ = λ {Y Z} {f g} eq → F.F-resp-≈ (second-cong eq)
}
where module F = Functor F

Expand All @@ -79,7 +79,7 @@ module FromCartesian o′ ℓ′ {o ℓ e} {C : Category o ℓ e} (Car : Cartesi
in begin
F₁ (first (f C.∘ g)) ⟨$⟩ x ≈˘⟨ [ F ]-resp-∘ first∘first ⟩
F₁ (first g) ⟨$⟩ (F₁ (first f) ⟨$⟩ x) ∎
; F-resp-≈ = λ {A B} {f g} eq → F-resp-≈ (-cong eq C.Equiv.refl)
; F-resp-≈ = λ {A B} {f g} eq → F-resp-≈ (first-cong eq)
}
where open Functor F

Expand Down Expand Up @@ -113,7 +113,7 @@ module FromCartesian o′ ℓ′ {o ℓ e} {C : Category o ℓ e} (Car : Cartesi
F.₁ C.id ⟨$⟩ (α.η Y ⟨$⟩ x) ≈⟨ F.identity ⟩
α.η Y ⟨$⟩ x ∎
; homomorphism = λ {X Y Z} → Setoid.sym (F.₀ (Z × _)) ([ F ]-resp-∘ first∘first)
; F-resp-≈ = λ eq → F.F-resp-≈ (-cong eq C.Equiv.refl)
; F-resp-≈ = λ eq → F.F-resp-≈ (first-cong eq)
}

module FromCartesianCCC {o} {C : Category o o o} (Car : Cartesian C) where
Expand Down
14 changes: 9 additions & 5 deletions src/Categories/Category/Construction/Pullbacks.agda
Original file line number Diff line number Diff line change
Expand Up @@ -33,12 +33,16 @@ record Pullback⇒ W (X Y : PullbackObj W) : Set (o ⊔ ℓ ⊔ e) where
commute₁ : Y.arr₁ ∘ mor₁ ≈ X.arr₁
commute₂ : Y.arr₂ ∘ mor₂ ≈ X.arr₂

private abstract
pbarrlemma : Y.arr₁ ∘ mor₁ ∘ X.pullback.p₁ ≈ Y.arr₂ ∘ mor₂ ∘ X.pullback.p₂
pbarrlemma = begin
Y.arr₁ ∘ mor₁ ∘ X.pullback.p₁ ≈⟨ pullˡ commute₁ ⟩
X.arr₁ ∘ X.pullback.p₁ ≈⟨ X.pullback.commute ⟩
X.arr₂ ∘ X.pullback.p₂ ≈˘⟨ pullˡ commute₂ ⟩
Y.arr₂ ∘ mor₂ ∘ X.pullback.p₂ ∎

pbarr : X.pullback.P ⇒ Y.pullback.P
pbarr = Y.pullback.universal $ begin
Y.arr₁ ∘ mor₁ ∘ X.pullback.p₁ ≈⟨ pullˡ commute₁ ⟩
X.arr₁ ∘ X.pullback.p₁ ≈⟨ X.pullback.commute ⟩
X.arr₂ ∘ X.pullback.p₂ ≈˘⟨ pullˡ commute₂ ⟩
Y.arr₂ ∘ mor₂ ∘ X.pullback.p₂ ∎
pbarr = Y.pullback.universal pbarrlemma

open Pullback⇒

Expand Down
48 changes: 6 additions & 42 deletions src/Categories/Functor/Slice.agda
Original file line number Diff line number Diff line change
Expand Up @@ -42,59 +42,24 @@ module _ {A : Obj} where

open S C

Forgetful : Functor (Slice A) C
Forgetful = record
TotalSpace : Functor (Slice A) C
TotalSpace = record
{ F₀ = Y
; F₁ = h
; identity = refl
; homomorphism = refl
; F-resp-≈ = id→
}

module _ (pullback : ∀ {X} {Y} {Z} (h : X ⇒ Z) (i : Y ⇒ Z) → Pullback h i) where
private
module pullbacks {X Y Z} h i = Pullback (pullback {X} {Y} {Z} h i)
open pullbacks using (p₂; p₂∘universal≈h₂; unique; unique-diagram; p₁∘universal≈h₁)

pullback-functorial : ∀ {B} (f : B ⇒ A) → Functor (Slice A) C
pullback-functorial f = record
{ F₀ = λ X → p.P X
; F₁ = λ f → p⇒ _ _ f
; identity = λ {X} → sym (p.unique X id-comm id-comm)
; homomorphism = λ {_ Y Z} →
p.unique-diagram Z (p.p₁∘universal≈h₁ Z ○ ⟺ identityˡ ○ ⟺ (pullʳ (p.p₁∘universal≈h₁ Y)) ○ ⟺ (pullˡ (p.p₁∘universal≈h₁ Z)))
(p.p₂∘universal≈h₂ Z ○ assoc ○ ⟺ (pullʳ (p.p₂∘universal≈h₂ Y)) ○ ⟺ (pullˡ (p.p₂∘universal≈h₂ Z)))
; F-resp-≈ = λ {_ B} {h i} eq →
p.unique-diagram B (p.p₁∘universal≈h₁ B ○ ⟺ (p.p₁∘universal≈h₁ B))
(p.p₂∘universal≈h₂ B ○ ∘-resp-≈ˡ eq ○ ⟺ (p.p₂∘universal≈h₂ B))
}
where p : ∀ X → Pullback f (arr X)
p X = pullback f (arr X)
module p X = Pullback (p X)

p⇒ : ∀ X Y (g : Slice⇒ X Y) → p.P X ⇒ p.P Y
p⇒ X Y g = Pbs.Pullback⇒.pbarr pX⇒pY
where pX : Pbs.PullbackObj C A
pX = record { pullback = p X }
pY : Pbs.PullbackObj C A
pY = record { pullback = p Y }
pX⇒pY : Pbs.Pullback⇒ C A pX pY
pX⇒pY = record
{ mor₁ = Category.id C
; mor₂ = h g
; commute₁ = identityʳ
; commute₂ = △ g
}

module _ (product : {X : Obj} → Product A X) where

private
module product {X} = Product (product {X})
open product

-- this is adapted from proposition 1.33 of Aspects of Topoi (Freyd, 1972)
Free : Functor C (Slice A)
Free = record
ConstantFamily : Functor C (Slice A)
ConstantFamily = record
{ F₀ = λ _ → sliceobj π₁
; F₁ = λ f → slicearr ([ product ⇒ product ]π₁∘× ○ identityˡ)
; identity = id×id product
Expand All @@ -112,9 +77,8 @@ module _ {A : Obj} where
open Terminal terminal
open BinaryProducts products

-- Needs better name!
Coforgetful : Functor (Slice A) C
Coforgetful = record
InternalSections : Functor (Slice A) C
InternalSections = record
{ F₀ = p.P
; F₁ = λ f → p.universal _ (F₁-lemma f)
; identity = λ {X} → sym $ p.unique X (sym (!-unique _)) $ begin
Expand Down
6 changes: 3 additions & 3 deletions src/Categories/Functor/Slice/BaseChange.agda
Original file line number Diff line number Diff line change
Expand Up @@ -7,7 +7,7 @@ module Categories.Functor.Slice.BaseChange {o ℓ e} (C : Category o ℓ e) wher
open import Categories.Category.Slice C using (Slice)
open import Categories.Category.Slice.Properties C using (pullback⇒product; slice-slice⇒slice; slice⇒slice-slice)
open import Categories.Functor using (Functor; _∘F_)
open import Categories.Functor.Slice using (Forgetful; Free)
open import Categories.Functor.Slice using (TotalSpace; ConstantFamily)
open import Categories.Diagram.Pullback C using (Pullback)

open Category C
Expand All @@ -16,9 +16,9 @@ module _ {A B : Obj} (f : B ⇒ A) where

-- Any morphism induces a functor between slices.
BaseChange! : Functor (Slice B) (Slice A)
BaseChange! = Forgetful (Slice A) ∘F slice⇒slice-slice f
BaseChange! = TotalSpace (Slice A) ∘F slice⇒slice-slice f

-- Any morphism which admits pullbacks induces a functor the other way between slices.
-- This is adjoint to BaseChange!: see Categories.Adjoint.Instance.BaseChange.
BaseChange* : (∀ {C} {h : C ⇒ A} → Pullback f h) → Functor (Slice A) (Slice B)
BaseChange* pullback = slice-slice⇒slice f ∘F Free (Slice A) (pullback⇒product pullback)
BaseChange* pullback = slice-slice⇒slice f ∘F ConstantFamily (Slice A) (pullback⇒product pullback)
Original file line number Diff line number Diff line change
Expand Up @@ -59,7 +59,7 @@ NNO×CCC⇒PNNO nno = record
commute₂' {A} {X} {f} {g} = begin
g ∘ (eval′ ∘ (universal (λg (f ∘ π₂)) (λg (g ∘ eval′)) ⁂ id)) ∘ swap ≈⟨ pullˡ (pullˡ (⟺ β′)) ⟩
((eval′ ∘ (λg (g ∘ eval′) ⁂ id)) ∘ (universal (λg (f ∘ π₂)) (λg (g ∘ eval′)) ⁂ id)) ∘ swap ≈⟨ (pullʳ ⁂∘⁂) ⟩∘⟨refl ⟩
(eval′ ∘ (λg (g ∘ eval′) ∘ universal (λg (f ∘ π₂)) (λg (g ∘ eval′)) ⁂ id ∘ id)) ∘ swap ≈⟨ (refl⟩∘⟨ (-cong s-commute refl)) ⟩∘⟨refl ⟩
(eval′ ∘ (λg (g ∘ eval′) ∘ universal (λg (f ∘ π₂)) (λg (g ∘ eval′)) ⁂ id ∘ id)) ∘ swap ≈⟨ (refl⟩∘⟨ (first-cong s-commute)) ⟩∘⟨refl ⟩
(eval′ ∘ (universal (λg (f ∘ π₂)) (λg (g ∘ eval′)) ∘ s ⁂ id ∘ id)) ∘ swap ≈⟨ (refl⟩∘⟨ (⟺ ⁂∘⁂)) ⟩∘⟨refl ⟩
(eval′ ∘ (universal (λg (f ∘ π₂)) (λg (g ∘ eval′)) ⁂ id) ∘ (s ⁂ id)) ∘ swap ≈⟨ pullʳ (pullʳ (⟺ swap∘⁂)) ⟩
eval′ ∘ (universal (λg (f ∘ π₂)) (λg (g ∘ eval′)) ⁂ id) ∘ swap ∘ (id ⁂ s) ≈⟨ sym-assoc ○ sym-assoc ⟩
Expand All @@ -71,7 +71,7 @@ NNO×CCC⇒PNNO nno = record
λg (u ∘ swap) ≈⟨ nno-unique
(⟺ (λ-unique′
(begin
eval′ ∘ (λg (u ∘ swap) ∘ z ⁂ id) ≈⟨ refl⟩∘⟨ ⟺ (⁂∘⁂ ○ -cong₂ refl identity²) ⟩
eval′ ∘ (λg (u ∘ swap) ∘ z ⁂ id) ≈⟨ refl⟩∘⟨ ⟺ (⁂∘⁂ ○ second-cong identity²) ⟩
eval′ ∘ (λg (u ∘ swap) ⁂ id) ∘ (z ⁂ id) ≈⟨ pullˡ β′ ⟩
(u ∘ swap) ∘ (z ⁂ id) ≈⟨ pullʳ swap∘⁂ ⟩
u ∘ (id ⁂ z) ∘ swap ≈⟨ pushʳ (bp-unique′
Expand Down
Loading