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 all 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)
106 changes: 102 additions & 4 deletions src/Categories/Adjoint/Instance/Slice.agda
Original file line number Diff line number Diff line change
Expand Up @@ -5,10 +5,18 @@ open import Categories.Category using (Category)
module Categories.Adjoint.Instance.Slice {o ℓ e} (C : Category o ℓ e) where

open import Categories.Adjoint using (_⊣_)
open import Categories.Category.Slice C using (SliceObj; Slice⇒; slicearr)
open import Categories.Functor.Slice C using (Forgetful; Free)
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 (TotalSpace; ConstantFamily; InternalSections)
open import Categories.Morphism.Reasoning C
open import Categories.NaturalTransformation using (ntHelper)
open import Categories.Diagram.Pullback C hiding (swap)
open import Categories.Object.Product C
open import Categories.Object.Terminal C

open import Function.Base using (_$_)

open Category C
open HomReasoning
Expand All @@ -22,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 All @@ -44,3 +52,93 @@ module _ {A : Obj} (product : ∀ {X} → Product A X) where
⟨ π₁ , π₂ ⟩ ≈⟨ η ⟩
id ∎
}

module _ {A : Obj} (ccc : CartesianClosed C) (pullback : ∀ {X} {Y} {Z} (h : X ⇒ Z) (i : Y ⇒ Z) → Pullback h i) where

open CartesianClosed ccc
open Cartesian cartesian
open Terminal terminal
open BinaryProducts products

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-≈ʳ (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 π₁)) ⟩
p.p₂ (sliceobj π₁) ∘ p.universal (sliceobj π₁) _ ∘ p.universal (sliceobj π₁) _ ∎
}
; counit = ntHelper record
{ η = λ X → slicearr (counit-△ 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⟩∘⟨ 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 ∎
}
; zig = λ {X} → begin
(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⟩∘⟨ first-cong (p.p₂∘universal≈h₂ (sliceobj π₁)) ⟩∘⟨refl ⟩
eval′ ∘ first (λg swap) ∘ swap ≈⟨ pullˡ β′ ⟩
swap ∘ swap ≈⟨ swap∘swap ⟩
id ∎
; zag = λ {X} → p.unique-diagram X !-unique₂ $ begin
p.p₂ X ∘ p.universal X _ ∘ p.universal (sliceobj π₁) _ ≈⟨ pullˡ (p.p₂∘universal≈h₂ X) ⟩
λg ((eval′ ∘ first (p.p₂ X) ∘ swap) ∘ eval′ ∘ first (p.p₂ (sliceobj π₁))) ∘ p.universal (sliceobj π₁) _ ≈˘⟨ pullˡ (subst ○ λ-cong assoc) ⟩
λ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ʳ (cancelʳ swap∘swap)) ⟩
λg (eval′ ∘ first (p.p₂ X)) ≈⟨ η-exp′ ⟩
p.p₂ X ≈˘⟨ identityʳ ⟩
p.p₂ X ∘ id ∎
}
where
p : (X : SliceObj A) → Pullback {X = ⊤} {Z = A ^ A} {Y = Y X ^ A} (λg π₂) (λg (arr X ∘ eval′))
p X = pullback (λg π₂) (λg (arr X ∘ eval′))
module p X = Pullback (p X)

abstract
unit-pb : ∀ (X : Obj) → eval′ ∘ first {A = X} {C = A} (λg π₂ ∘ !) ≈ eval′ ∘ first (λg (π₁ ∘ eval′) ∘ λg swap)
unit-pb X = begin
eval′ ∘ first (λg π₂ ∘ !) ≈˘⟨ refl⟩∘⟨ first∘first ⟩
eval′ ∘ first (λg π₂) ∘ first ! ≈⟨ pullˡ β′ ⟩
π₂ ∘ first ! ≈⟨ π₂∘⁂ ○ identityˡ ⟩
π₂ ≈˘⟨ project₁ ⟩
π₁ ∘ swap ≈˘⟨ refl⟩∘⟨ β′ ⟩
π₁ ∘ eval′ ∘ first (λg swap) ≈˘⟨ extendʳ β′ ⟩
eval′ ∘ first (λg (π₁ ∘ eval′)) ∘ first (λg swap) ≈⟨ refl⟩∘⟨ first∘first ⟩
eval′ ∘ first (λg (π₁ ∘ eval′) ∘ λg swap) ∎
-- A good chunk of the above, maybe all if you squint, is duplicated with F₁-lemma
-- eval′ ∘ first (λg π₂ ∘ !) ≈ eval′ ∘ first (λg (f ∘ eval′) ∘ first (λg g)
-- With f : X ⇒ Y and g : Z × Y ⇒ X. Not sure what conditions f and g need to have

-- Would it be better if Free used π₂ rather than π₁?
-- It would mean we could avoid this swap
counit-△ : ∀ X → arr X ∘ eval′ ∘ first (p.p₂ X) ∘ swap ≈ π₁
counit-△ X = begin
arr X ∘ eval′ ∘ first (p.p₂ X) ∘ swap ≈˘⟨ assoc² ⟩
((arr X ∘ eval′) ∘ first (p.p₂ X)) ∘ swap ≈⟨ lemma ⟩∘⟨refl ⟩
(π₂ ∘ first (p.p₁ X)) ∘ swap ≈⟨ (π₂∘⁂ ○ identityˡ) ⟩∘⟨refl ⟩
π₂ ∘ swap ≈⟨ project₂ ⟩
π₁ ∎
where
lemma : (arr X ∘ eval′) ∘ first (p.p₂ X) ≈ π₂ ∘ first (p.p₁ X)
lemma = λ-inj $ begin
λg ((arr X ∘ eval′) ∘ first (p.p₂ X)) ≈˘⟨ subst ⟩
λg (arr X ∘ eval′) ∘ p.p₂ X ≈˘⟨ p.commute X ⟩
λg π₂ ∘ p.p₁ X ≈⟨ subst ⟩
λg (π₂ ∘ first (p.p₁ X)) ∎


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
Loading
Loading