cat/src/Cat/Categories/Cat.agda

365 lines
15 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

{-# OPTIONS --cubical --allow-unsolved-metas #-}
module Cat.Categories.Cat where
open import Agda.Primitive
open import Cubical
open import Function
open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
open import Cat.Category
open import Cat.Functor
open import Cat.Equality
open Equality.Data.Product
open Functor
open IsFunctor
open Category hiding (_∘_)
-- The category of categories
module _ ( ' : Level) where
private
module _ {𝔸 𝔹 𝔻 : Category '} {F : Functor 𝔸 𝔹} {G : Functor 𝔹 } {H : Functor 𝔻} where
private
eq* : func* (H ∘f (G ∘f F)) func* ((H ∘f G) ∘f F)
eq* = refl
eq→ : PathP
(λ i {A B : 𝔸 .Object} 𝔸 [ A , B ] 𝔻 [ eq* i A , eq* i B ])
(func→ (H ∘f (G ∘f F))) (func→ ((H ∘f G) ∘f F))
eq→ = refl
postulate
eqI
: (λ i {A : 𝔸 .Object} eq→ i (𝔸 .𝟙 {A}) 𝔻 .𝟙 {eq* i A})
[ (H ∘f (G ∘f F)) .isFunctor .ident
((H ∘f G) ∘f F) .isFunctor .ident
]
eqD
: (λ i {A B C} {f : 𝔸 [ A , B ]} {g : 𝔸 [ B , C ]}
eq→ i (𝔸 [ g f ]) 𝔻 [ eq→ i g eq→ i f ])
[ (H ∘f (G ∘f F)) .isFunctor .distrib
((H ∘f G) ∘f F) .isFunctor .distrib
]
assc : H ∘f (G ∘f F) (H ∘f G) ∘f F
assc = Functor≡ eq* eq→ (IsFunctor≡ eqI eqD)
module _ { 𝔻 : Category '} {F : Functor 𝔻} where
module _ where
private
eq* : (func* F) (func* (identity {C = })) func* F
eq* = refl
-- lemmm : func→ {C = A} {D = B} (f ∘f identity) ≡ func→ f
eq→ : PathP
(λ i
{x y : Object } Arrow x y Arrow 𝔻 (func* F x) (func* F y))
(func→ (F ∘f identity)) (func→ F)
eq→ = refl
postulate
eqI-r
: (λ i {c : .Object} (λ _ 𝔻 [ func* F c , func* F c ])
[ func→ F ( .𝟙) 𝔻 .𝟙 ])
[(F ∘f identity) .isFunctor .ident F .isFunctor .ident ]
eqD-r : PathP
(λ i
{A B C : .Object} {f : .Arrow A B} {g : .Arrow B C}
eq→ i ( [ g f ]) 𝔻 [ eq→ i g eq→ i f ])
((F ∘f identity) .isFunctor .distrib) (F .isFunctor .distrib)
ident-r : F ∘f identity F
ident-r = Functor≡ eq* eq→ (IsFunctor≡ eqI-r eqD-r)
module _ where
private
postulate
eq* : (identity ∘f F) .func* F .func*
eq→ : PathP
(λ i {x y : Object } .Arrow x y 𝔻 .Arrow (eq* i x) (eq* i y))
((identity ∘f F) .func→) (F .func→)
eqI : (λ i {A : .Object} eq→ i ( .𝟙 {A}) 𝔻 .𝟙 {eq* i A})
[ ((identity ∘f F) .isFunctor .ident) (F .isFunctor .ident) ]
eqD : PathP (λ i {A B C : .Object} {f : .Arrow A B} {g : .Arrow B C}
eq→ i ( [ g f ]) 𝔻 [ eq→ i g eq→ i f ])
((identity ∘f F) .isFunctor .distrib) (F .isFunctor .distrib)
-- (λ z → eq* i z) (eq→ i)
ident-l : identity ∘f F F
ident-l = Functor≡ eq* eq→ λ i record { ident = eqI i ; distrib = eqD i }
Cat : Category (lsuc ( ')) ( ')
Cat =
record
{ Object = Category '
; Arrow = Functor
; 𝟙 = identity
; _∘_ = _∘f_
-- What gives here? Why can I not name the variables directly?
; isCategory = record
{ assoc = λ {_ _ _ _ F G H} assc {F = F} {G = G} {H = H}
; ident = ident-r , ident-l
}
}
module _ { ' : Level} where
module _ ( 𝔻 : Category ') where
private
Catt = Cat '
:Object: = .Object × 𝔻 .Object
:Arrow: : :Object: :Object: Set '
:Arrow: (c , d) (c' , d') = Arrow c c' × Arrow 𝔻 d d'
:𝟙: : {o : :Object:} :Arrow: o o
:𝟙: = .𝟙 , 𝔻 .𝟙
_:⊕:_ :
{a b c : :Object:}
:Arrow: b c
:Arrow: a b
:Arrow: a c
_:⊕:_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) [ bc∈C ab∈C ] , 𝔻 [ bc∈D ab∈D ]}
instance
:isCategory: : IsCategory :Object: :Arrow: :𝟙: _:⊕:_
:isCategory: = record
{ assoc = Σ≡ C.assoc D.assoc
; ident
= Σ≡ (fst C.ident) (fst D.ident)
, Σ≡ (snd C.ident) (snd D.ident)
}
where
open module C = IsCategory ( .isCategory)
open module D = IsCategory (𝔻 .isCategory)
:product: : Category '
:product: = record
{ Object = :Object:
; Arrow = :Arrow:
; 𝟙 = :𝟙:
; _∘_ = _:⊕:_
}
proj₁ : Catt [ :product: , ]
proj₁ = record { func* = fst ; func→ = fst ; isFunctor = record { ident = refl ; distrib = refl } }
proj₂ : Catt [ :product: , 𝔻 ]
proj₂ = record { func* = snd ; func→ = snd ; isFunctor = record { ident = refl ; distrib = refl } }
module _ {X : Object Catt} (x₁ : Catt [ X , ]) (x₂ : Catt [ X , 𝔻 ]) where
open Functor
x : Functor X :product:
x = record
{ func* = λ x x₁ .func* x , x₂ .func* x
; func→ = λ x func→ x₁ x , func→ x₂ x
; isFunctor = record
{ ident = Σ≡ x₁.ident x₂.ident
; distrib = Σ≡ x₁.distrib x₂.distrib
}
}
where
open module x = IsFunctor (x₁ .isFunctor)
open module x = IsFunctor (x₂ .isFunctor)
isUniqL : Catt [ proj₁ x ] x₁
isUniqL = Functor≡ eq* eq→ eqIsF -- Functor≡ refl refl λ i → record { ident = refl i ; distrib = refl i }
where
eq* : (Catt [ proj₁ x ]) .func* x₁ .func*
eq* = refl
eq→ : (λ i {A : X .Object} {B : X .Object} X [ A , B ] [ eq* i A , eq* i B ])
[ (Catt [ proj₁ x ]) .func→ x₁ .func→ ]
eq→ = refl
postulate eqIsF : (Catt [ proj₁ x ]) .isFunctor x₁ .isFunctor
-- eqIsF = IsFunctor≡ {!refl!} {!!}
postulate isUniqR : Catt [ proj₂ x ] x₂
-- isUniqR = Functor≡ refl refl {!!} {!!}
isUniq : Catt [ proj₁ x ] x₁ × Catt [ proj₂ x ] x₂
isUniq = isUniqL , isUniqR
uniq : ∃![ x ] (Catt [ proj₁ x ] x₁ × Catt [ proj₂ x ] x₂)
uniq = x , isUniq
instance
isProduct : IsProduct (Cat ') proj₁ proj₂
isProduct = uniq
product : Product { = (Cat ')} 𝔻
product = record
{ obj = :product:
; proj₁ = proj₁
; proj₂ = proj₂
}
module _ { ' : Level} where
instance
hasProducts : HasProducts (Cat ')
hasProducts = record { product = product }
-- Basically proves that `Cat ` is cartesian closed.
module _ ( : Level) where
private
open Data.Product
open import Cat.Categories.Fun
Cat : Category (lsuc ( )) ( )
Cat = Cat
module _ ( 𝔻 : Category ) where
private
:obj: : Cat .Object
:obj: = Fun { = } {𝔻 = 𝔻}
:func*: : Functor 𝔻 × .Object 𝔻 .Object
:func*: (F , A) = F .func* A
module _ {dom cod : Functor 𝔻 × .Object} where
private
F : Functor 𝔻
F = proj₁ dom
A : .Object
A = proj₂ dom
G : Functor 𝔻
G = proj₁ cod
B : .Object
B = proj₂ cod
:func→: : (pobj : NaturalTransformation F G × .Arrow A B)
𝔻 .Arrow (F .func* A) (G .func* B)
:func→: ((θ , θNat) , f) = result
where
θA : 𝔻 .Arrow (F .func* A) (G .func* A)
θA = θ A
θB : 𝔻 .Arrow (F .func* B) (G .func* B)
θB = θ B
F→f : 𝔻 .Arrow (F .func* A) (F .func* B)
F→f = F .func→ f
G→f : 𝔻 .Arrow (G .func* A) (G .func* B)
G→f = G .func→ f
l : 𝔻 .Arrow (F .func* A) (G .func* B)
l = 𝔻 [ θB F→f ]
r : 𝔻 .Arrow (F .func* A) (G .func* B)
r = 𝔻 [ G→f θA ]
-- There are two choices at this point,
-- but I suppose the whole point is that
-- by `θNat f` we have `l ≡ r`
-- lem : 𝔻 [ θ B ∘ F .func→ f ] ≡ 𝔻 [ G .func→ f ∘ θ A ]
-- lem = θNat f
result : 𝔻 .Arrow (F .func* A) (G .func* B)
result = l
_×p_ = product
module _ {c : Functor 𝔻 × .Object} where
private
F : Functor 𝔻
F = proj₁ c
C : .Object
C = proj₂ c
-- NaturalTransformation F G × .Arrow A B
-- :ident: : :func→: {c} {c} (identityNat F , .𝟙) 𝔻 .𝟙
-- :ident: = trans (proj₂ 𝔻.ident) (F .ident)
-- where
-- open module 𝔻 = IsCategory (𝔻 .isCategory)
-- Unfortunately the equational version has some ambigous arguments.
:ident: : :func→: {c} {c} (identityNat F , .𝟙 {o = proj₂ c}) 𝔻 .𝟙
:ident: = begin
:func→: {c} {c} ((:obj: ×p ) .Product.obj .𝟙 {c}) ≡⟨⟩
:func→: {c} {c} (identityNat F , .𝟙) ≡⟨⟩
𝔻 [ identityTrans F C F .func→ ( .𝟙)] ≡⟨⟩
𝔻 [ 𝔻 .𝟙 F .func→ ( .𝟙)] ≡⟨ proj₂ 𝔻.ident
F .func→ ( .𝟙) ≡⟨ F.ident
𝔻 .𝟙
where
open module 𝔻 = IsCategory (𝔻 .isCategory)
open module F = IsFunctor (F .isFunctor)
module _ {F×A G×B H×C : Functor 𝔻 × .Object} where
F = F×A .proj₁
A = F×A .proj₂
G = G×B .proj₁
B = G×B .proj₂
H = H×C .proj₁
C = H×C .proj₂
-- Not entirely clear what this is at this point:
_P⊕_ = (:obj: ×p ) .Product.obj .Category._∘_ {F×A} {G×B} {H×C}
module _
-- NaturalTransformation F G × .Arrow A B
{θ×f : NaturalTransformation F G × .Arrow A B}
{η×g : NaturalTransformation G H × .Arrow B C} where
private
θ : Transformation F G
θ = proj₁ (proj₁ θ×f)
θNat : Natural F G θ
θNat = proj₂ (proj₁ θ×f)
f : .Arrow A B
f = proj₂ θ×f
η : Transformation G H
η = proj₁ (proj₁ η×g)
ηNat : Natural G H η
ηNat = proj₂ (proj₁ η×g)
g : .Arrow B C
g = proj₂ η×g
ηθNT : NaturalTransformation F H
ηθNT = Fun .Category._∘_ {F} {G} {H} (η , ηNat) (θ , θNat)
ηθ = proj₁ ηθNT
ηθNat = proj₂ ηθNT
:distrib: :
𝔻 [ 𝔻 [ η C θ C ] F .func→ ( [ g f ] ) ]
𝔻 [ 𝔻 [ η C G .func→ g ] 𝔻 [ θ B F .func→ f ] ]
:distrib: = begin
𝔻 [ (ηθ C) F .func→ ( [ g f ]) ]
≡⟨ ηθNat ( [ g f ])
𝔻 [ H .func→ ( [ g f ]) (ηθ A) ]
≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.distrib)
𝔻 [ 𝔻 [ H .func→ g H .func→ f ] (ηθ A) ]
≡⟨ sym assoc
𝔻 [ H .func→ g 𝔻 [ H .func→ f ηθ A ] ]
≡⟨ cong (λ φ 𝔻 [ H .func→ g φ ]) assoc
𝔻 [ H .func→ g 𝔻 [ 𝔻 [ H .func→ f η A ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ H .func→ g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f)))
𝔻 [ H .func→ g 𝔻 [ 𝔻 [ η B G .func→ f ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ H .func→ g φ ]) (sym assoc)
𝔻 [ H .func→ g 𝔻 [ η B 𝔻 [ G .func→ f θ A ] ] ] ≡⟨ assoc
𝔻 [ 𝔻 [ H .func→ g η B ] 𝔻 [ G .func→ f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ G .func→ f θ A ] ]) (sym (ηNat g))
𝔻 [ 𝔻 [ η C G .func→ g ] 𝔻 [ G .func→ f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C G .func→ g ] φ ]) (sym (θNat f))
𝔻 [ 𝔻 [ η C G .func→ g ] 𝔻 [ θ B F .func→ f ] ]
where
open IsCategory (𝔻 .isCategory)
open module H = IsFunctor (H .isFunctor)
:eval: : Functor ((:obj: ×p ) .Product.obj) 𝔻
:eval: = record
{ func* = :func*:
; func→ = λ {dom} {cod} :func→: {dom} {cod}
; isFunctor = record
{ ident = λ {o} :ident: {o}
; distrib = λ {f u n k y} :distrib: {f} {u} {n} {k} {y}
}
}
module _ (𝔸 : Category ) (F : Functor ((𝔸 ×p ) .Product.obj) 𝔻) where
open HasProducts (hasProducts {} {}) using (parallelProduct)
postulate
transpose : Functor 𝔸 :obj:
eq : Cat [ :eval: (parallelProduct transpose (Cat .𝟙 {o = })) ] F
catTranspose : ∃![ F~ ] (Cat [ :eval: (parallelProduct F~ (Cat .𝟙 {o = }))] F )
catTranspose = transpose , eq
:isExponential: : IsExponential Cat 𝔻 :obj: :eval:
:isExponential: = catTranspose
-- :exponent: : Exponential (Cat ) A B
:exponent: : Exponential Cat 𝔻
:exponent: = record
{ obj = :obj:
; eval = :eval:
; isExponential = :isExponential:
}
hasExponentials : HasExponentials (Cat )
hasExponentials = record { exponent = :exponent: }