cat/src/Cat/Categories/Cat.agda

366 lines
15 KiB
Agda
Raw Normal View History

2018-02-02 14:33:54 +00:00
-- There is no category of categories in our interpretation
{-# 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
2018-01-31 13:39:54 +00:00
open import Cat.Equality
open Equality.Data.Product
open Functor
2018-01-30 15:23:36 +00:00
open IsFunctor
open Category hiding (_∘_)
2018-01-25 11:01:37 +00:00
-- The category of categories
2018-01-24 15:38:28 +00:00
module _ ( ' : Level) where
private
2018-01-31 13:39:54 +00:00
module _ {𝔸 𝔹 𝔻 : Category '} {F : Functor 𝔸 𝔹} {G : Functor 𝔹 } {H : Functor 𝔻} where
2018-01-25 11:44:47 +00:00
private
2018-01-31 13:39:54 +00:00
eq* : func* (H ∘f (G ∘f F)) func* ((H ∘f G) ∘f F)
2018-01-25 11:44:47 +00:00
eq* = refl
eq→ : PathP
2018-01-31 13:39:54 +00:00
(λ i {A B : 𝔸 .Object} 𝔸 [ A , B ] 𝔻 [ eq* i A , eq* i B ])
(func→ (H ∘f (G ∘f F))) (func→ ((H ∘f G) ∘f F))
2018-01-25 11:44:47 +00:00
eq→ = refl
2018-01-31 13:39:54 +00:00
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)
2018-01-25 11:44:47 +00:00
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 ]
2018-01-25 11:44:47 +00:00
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 ])
2018-01-30 15:23:36 +00:00
((F ∘f identity) .isFunctor .distrib) (F .isFunctor .distrib)
2018-01-25 11:44:47 +00:00
ident-r : F ∘f identity F
ident-r = Functor≡ eq* eq→ (IsFunctor≡ eqI-r eqD-r)
2018-01-25 11:44:47 +00:00
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) ]
2018-01-25 11:44:47 +00:00
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 ])
2018-01-30 15:23:36 +00:00
((identity ∘f F) .isFunctor .distrib) (F .isFunctor .distrib)
-- (λ z → eq* i z) (eq→ i)
2018-01-25 11:44:47 +00:00
ident-l : identity ∘f F F
ident-l = Functor≡ eq* eq→ λ i record { ident = eqI i ; distrib = eqD i }
2018-01-21 14:23:40 +00:00
Cat : Category (lsuc ( ')) ( ')
Cat =
record
{ Object = Category '
; Arrow = Functor
; 𝟙 = identity
; _∘_ = _∘f_
-- What gives here? Why can I not name the variables directly?
2018-01-21 18:23:24 +00:00
; isCategory = record
2018-01-31 13:39:54 +00:00
{ assoc = λ {_ _ _ _ F G H} assc {F = F} {G = G} {H = H}
2018-01-21 18:23:24 +00:00
; ident = ident-r , ident-l
}
}
2018-01-24 15:38:28 +00:00
module _ { ' : Level} where
2018-01-25 11:44:47 +00:00
module _ ( 𝔻 : Category ') where
2018-01-24 15:38:28 +00:00
private
2018-01-31 13:39:54 +00:00
Catt = Cat '
2018-01-25 11:44:47 +00:00
:Object: = .Object × 𝔻 .Object
2018-01-24 15:38:28 +00:00
:Arrow: : :Object: :Object: Set '
2018-01-25 11:44:47 +00:00
:Arrow: (c , d) (c' , d') = Arrow c c' × Arrow 𝔻 d d'
2018-01-24 15:38:28 +00:00
:𝟙: : {o : :Object:} :Arrow: o o
2018-01-25 11:44:47 +00:00
:𝟙: = .𝟙 , 𝔻 .𝟙
2018-01-24 15:38:28 +00:00
_:⊕:_ :
{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 ]}
2018-01-24 15:38:28 +00:00
instance
:isCategory: : IsCategory :Object: :Arrow: :𝟙: _:⊕:_
:isCategory: = record
2018-01-31 13:39:54 +00:00
{ assoc = Σ≡ C.assoc D.assoc
2018-01-24 15:38:28 +00:00
; ident
2018-01-31 13:39:54 +00:00
= Σ≡ (fst C.ident) (fst D.ident)
, Σ≡ (snd C.ident) (snd D.ident)
2018-01-24 15:38:28 +00:00
}
where
2018-01-25 11:44:47 +00:00
open module C = IsCategory ( .isCategory)
open module D = IsCategory (𝔻 .isCategory)
2018-01-24 15:38:28 +00:00
:product: : Category '
:product: = record
{ Object = :Object:
; Arrow = :Arrow:
; 𝟙 = :𝟙:
; _∘_ = _:⊕:_
}
2018-01-24 15:38:28 +00:00
2018-01-31 13:39:54 +00:00
proj₁ : Catt [ :product: , ]
2018-01-30 15:23:36 +00:00
proj₁ = record { func* = fst ; func→ = fst ; isFunctor = record { ident = refl ; distrib = refl } }
2018-01-24 15:38:28 +00:00
2018-01-31 13:39:54 +00:00
proj₂ : Catt [ :product: , 𝔻 ]
2018-01-30 15:23:36 +00:00
proj₂ = record { func* = snd ; func→ = snd ; isFunctor = record { ident = refl ; distrib = refl } }
2018-01-24 15:38:28 +00:00
2018-01-31 13:39:54 +00:00
module _ {X : Object Catt} (x₁ : Catt [ X , ]) (x₂ : Catt [ X , 𝔻 ]) where
2018-01-24 15:38:28 +00:00
open Functor
x : Functor X :product:
x = record
2018-01-31 13:39:54 +00:00
{ func* = λ x x₁ .func* x , x₂ .func* x
2018-01-24 15:38:28 +00:00
; func→ = λ x func→ x₁ x , func→ x₂ x
2018-01-30 15:23:36 +00:00
; isFunctor = record
2018-01-31 13:39:54 +00:00
{ ident = Σ≡ x₁.ident x₂.ident
; distrib = Σ≡ x₁.distrib x₂.distrib
2018-01-30 15:23:36 +00:00
}
2018-01-24 15:38:28 +00:00
}
2018-01-30 15:23:36 +00:00
where
open module x = IsFunctor (x₁ .isFunctor)
open module x = IsFunctor (x₂ .isFunctor)
2018-01-24 15:38:28 +00:00
2018-01-31 13:39:54 +00:00
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!} {!!}
2018-01-24 15:38:28 +00:00
postulate isUniqR : Catt [ proj₂ x ] x₂
2018-01-25 11:44:47 +00:00
-- isUniqR = Functor≡ refl refl {!!} {!!}
2018-01-24 15:38:28 +00:00
isUniq : Catt [ proj₁ x ] x₁ × Catt [ proj₂ x ] x₂
2018-01-24 15:38:28 +00:00
isUniq = isUniqL , isUniqR
uniq : ∃![ x ] (Catt [ proj₁ x ] x₁ × Catt [ proj₂ x ] x₂)
2018-01-24 15:38:28 +00:00
uniq = x , isUniq
2018-01-25 11:44:47 +00:00
instance
isProduct : IsProduct (Cat ') proj₁ proj₂
isProduct = uniq
2018-01-24 15:38:28 +00:00
2018-01-25 11:44:47 +00:00
product : Product { = (Cat ')} 𝔻
2018-01-24 15:38:28 +00:00
product = record
{ obj = :product:
; proj₁ = proj₁
; proj₂ = proj₂
}
2018-01-24 15:38:28 +00:00
module _ { ' : Level} where
instance
2018-01-25 11:01:37 +00:00
hasProducts : HasProducts (Cat ')
hasProducts = record { product = product }
2018-01-24 15:38:28 +00:00
-- Basically proves that `Cat ` is cartesian closed.
2018-01-25 11:01:37 +00:00
module _ ( : Level) where
2018-01-24 15:38:28 +00:00
private
2018-01-25 11:01:37 +00:00
open Data.Product
2018-01-24 15:38:28 +00:00
open import Cat.Categories.Fun
2018-01-25 11:01:37 +00:00
Cat : Category (lsuc ( )) ( )
Cat = Cat
2018-01-24 15:38:28 +00:00
module _ ( 𝔻 : Category ) where
private
:obj: : Cat .Object
:obj: = Fun { = } {𝔻 = 𝔻}
2018-01-24 15:38:28 +00:00
:func*: : Functor 𝔻 × .Object 𝔻 .Object
:func*: (F , A) = F .func* A
2018-01-24 15:38:28 +00:00
module _ {dom cod : Functor 𝔻 × .Object} where
private
F : Functor 𝔻
F = proj₁ dom
A : .Object
A = proj₂ dom
2018-01-24 15:38:28 +00:00
G : Functor 𝔻
G = proj₁ cod
B : .Object
B = proj₂ cod
2018-01-24 15:38:28 +00:00
: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 ]
2018-01-24 15:38:28 +00:00
r : 𝔻 .Arrow (F .func* A) (G .func* B)
r = 𝔻 [ G→f θA ]
2018-01-24 15:38:28 +00:00
-- 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 ]
2018-01-24 15:38:28 +00:00
-- lem = θNat f
result : 𝔻 .Arrow (F .func* A) (G .func* B)
result = l
2018-01-24 15:38:28 +00:00
_×p_ = product
2018-01-24 15:38:28 +00:00
module _ {c : Functor 𝔻 × .Object} where
private
F : Functor 𝔻
F = proj₁ c
C : .Object
C = proj₂ c
2018-01-24 15:38:28 +00:00
-- NaturalTransformation F G × .Arrow A B
2018-01-25 11:01:37 +00:00
-- :ident: : :func→: {c} {c} (identityNat F , .𝟙) 𝔻 .𝟙
-- :ident: = trans (proj₂ 𝔻.ident) (F .ident)
2018-01-24 15:38:28 +00:00
-- where
-- open module 𝔻 = IsCategory (𝔻 .isCategory)
2018-01-25 11:01:37 +00:00
-- 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
2018-01-30 15:23:36 +00:00
F .func→ ( .𝟙) ≡⟨ F.ident
𝔻 .𝟙
2018-01-25 11:01:37 +00:00
where
open module 𝔻 = IsCategory (𝔻 .isCategory)
2018-01-30 15:23:36 +00:00
open module F = IsFunctor (F .isFunctor)
2018-01-24 15:38:28 +00:00
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}
2018-01-24 15:38:28 +00:00
module _
-- NaturalTransformation F G × .Arrow A B
2018-01-25 11:01:37 +00:00
{θ×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)
2018-01-25 11:01:37 +00:00
ηθ = proj₁ ηθNT
ηθNat = proj₂ ηθNT
2018-01-24 15:38:28 +00:00
:distrib: :
𝔻 [ 𝔻 [ η C θ C ] F .func→ ( [ g f ] ) ]
𝔻 [ 𝔻 [ η C G .func→ g ] 𝔻 [ θ B F .func→ f ] ]
2018-01-24 15:38:28 +00:00
: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 ] ]
2018-01-24 15:38:28 +00:00
where
2018-01-25 11:01:37 +00:00
open IsCategory (𝔻 .isCategory)
2018-01-30 15:23:36 +00:00
open module H = IsFunctor (H .isFunctor)
2018-01-25 11:01:37 +00:00
2018-01-24 15:38:28 +00:00
:eval: : Functor ((:obj: ×p ) .Product.obj) 𝔻
:eval: = record
{ func* = :func*:
; func→ = λ {dom} {cod} :func→: {dom} {cod}
2018-01-30 15:23:36 +00:00
; isFunctor = record
{ ident = λ {o} :ident: {o}
; distrib = λ {f u n k y} :distrib: {f} {u} {n} {k} {y}
}
2018-01-24 15:38:28 +00:00
}
module _ (𝔸 : Category ) (F : Functor ((𝔸 ×p ) .Product.obj) 𝔻) where
2018-01-25 11:01:37 +00:00
open HasProducts (hasProducts {} {}) using (parallelProduct)
2018-01-24 15:38:28 +00:00
postulate
transpose : Functor 𝔸 :obj:
eq : Cat [ :eval: (parallelProduct transpose (Cat .𝟙 {o = })) ] F
2018-01-24 15:38:28 +00:00
catTranspose : ∃![ F~ ] (Cat [ :eval: (parallelProduct F~ (Cat .𝟙 {o = }))] F )
2018-01-24 15:38:28 +00:00
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:
}
2018-01-25 11:01:37 +00:00
hasExponentials : HasExponentials (Cat )
hasExponentials = record { exponent = :exponent: }