Prove that Cat is cartesian closed

WIP
This commit is contained in:
Frederik Hanghøj Iversen 2018-01-24 16:38:28 +01:00
parent 6a25a4c3ff
commit c5a3673d9b
5 changed files with 359 additions and 102 deletions

View file

@ -40,7 +40,7 @@ module _ { ' : Level} {A B : Category '} where
lift-eq-functors eq* eq→ eqI eqD i = record { func* = eq* i ; func→ = eq→ i ; ident = eqI i ; distrib = eqD i }
-- The category of categories
module _ { ' : Level} where
module _ ( ' : Level) where
private
module _ {A B C D : Category '} {f : Functor A B} {g : Functor B C} {h : Functor C D} where
eq* : func* (h ∘f (g ∘f f)) func* ((h ∘f g) ∘f f)
@ -94,81 +94,280 @@ module _ { ' : Level} where
}
}
module _ { : Level} (C D : Category ) where
private
:Object: = C .Object × D .Object
:Arrow: : :Object: :Object: Set
:Arrow: (c , d) (c' , d') = Arrow C c c' × Arrow D d d'
:𝟙: : {o : :Object:} :Arrow: o o
:𝟙: = C .𝟙 , D .𝟙
_:⊕:_ :
{a b c : :Object:}
:Arrow: b c
:Arrow: a b
:Arrow: a c
_:⊕:_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) (C ._⊕_) bc∈C ab∈C , D ._⊕_ bc∈D ab∈D}
module _ { ' : Level} where
Catt = Cat '
instance
:isCategory: : IsCategory :Object: :Arrow: :𝟙: _:⊕:_
:isCategory: = record
{ assoc = eqpair C.assoc D.assoc
; ident
= eqpair (fst C.ident) (fst D.ident)
, eqpair (snd C.ident) (snd D.ident)
module _ (C D : Category ') where
private
:Object: = C .Object × D .Object
:Arrow: : :Object: :Object: Set '
:Arrow: (c , d) (c' , d') = Arrow C c c' × Arrow D d d'
:𝟙: : {o : :Object:} :Arrow: o o
:𝟙: = C .𝟙 , D .𝟙
_:⊕:_ :
{a b c : :Object:}
:Arrow: b c
:Arrow: a b
:Arrow: a c
_:⊕:_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) (C ._⊕_) bc∈C ab∈C , D ._⊕_ bc∈D ab∈D}
instance
:isCategory: : IsCategory :Object: :Arrow: :𝟙: _:⊕:_
:isCategory: = record
{ assoc = eqpair C.assoc D.assoc
; ident
= eqpair (fst C.ident) (fst D.ident)
, eqpair (snd C.ident) (snd D.ident)
}
where
open module C = IsCategory (C .isCategory)
open module D = IsCategory (D .isCategory)
:product: : Category '
:product: = record
{ Object = :Object:
; Arrow = :Arrow:
; 𝟙 = :𝟙:
; _⊕_ = _:⊕:_
}
where
open module C = IsCategory (C .isCategory)
open module D = IsCategory (D .isCategory)
:product: : Category
:product: = record
{ Object = :Object:
; Arrow = :Arrow:
; 𝟙 = :𝟙:
; _⊕_ = _:⊕:_
proj₁ : Arrow Catt :product: C
proj₁ = record { func* = fst ; func→ = fst ; ident = refl ; distrib = refl }
proj₂ : Arrow Catt :product: D
proj₂ = record { func* = snd ; func→ = snd ; ident = refl ; distrib = refl }
module _ {X : Object Catt} (x₁ : Arrow Catt X C) (x₂ : Arrow Catt X D) where
open Functor
-- ident' : {c : Object X} → ((func→ x₁) {dom = c} (𝟙 X) , (func→ x₂) {dom = c} (𝟙 X)) ≡ 𝟙 (catProduct C D)
-- ident' {c = c} = lift-eq (ident x₁) (ident x₂)
x : Functor X :product:
x = record
{ func* = λ x (func* x₁) x , (func* x₂) x
; func→ = λ x func→ x₁ x , func→ x₂ x
; ident = lift-eq (ident x₁) (ident x₂)
; distrib = lift-eq (distrib x₁) (distrib x₂)
}
-- Need to "lift equality of functors"
-- If I want to do this like I do it for pairs it's gonna be a pain.
postulate isUniqL : (Catt proj₁) x x₁
-- isUniqL = lift-eq-functors refl refl {!!} {!!}
postulate isUniqR : (Catt proj₂) x x₂
-- isUniqR = lift-eq-functors 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 Catt proj₁ proj₂
isProduct = uniq
product : Product { = Catt} C D
product = record
{ obj = :product:
; proj₁ = proj₁
; proj₂ = proj₂
}
proj₁ : Arrow Cat :product: C
proj₁ = record { func* = fst ; func→ = fst ; ident = refl ; distrib = refl }
module _ { ' : Level} where
open Category
instance
CatHasProducts : HasProducts (Cat ')
CatHasProducts = record { product = product }
proj₂ : Arrow Cat :product: D
proj₂ = record { func* = snd ; func→ = snd ; ident = refl ; distrib = refl }
-- Basically proves that `Cat ` is cartesian closed.
module _ { : Level} { : Category } {{_ : HasProducts (Opposite )}} where
open Data.Product
open Category
module _ {X : Object (Cat {} {})} (x₁ : Arrow Cat X C) (x₂ : Arrow Cat X D) where
open Functor
private
Cat : Category (lsuc ( )) ( )
Cat = Cat
open import Cat.Categories.Fun
open Functor
module _ ( 𝔻 : Category ) where
private
:obj: : Cat .Object
:obj: = Fun { = } {𝔻 = 𝔻}
-- ident' : {c : Object X} → ((func→ x₁) {dom = c} (𝟙 X) , (func→ x₂) {dom = c} (𝟙 X)) ≡ 𝟙 (catProduct C D)
-- ident' {c = c} = lift-eq (ident x₁) (ident x₂)
:func*: : Functor 𝔻 × .Object 𝔻 .Object
:func*: (F , A) = F .func* A
x : Functor X :product:
x = record
{ func* = λ x (func* x₁) x , (func* x₂) x
; func→ = λ x func→ x₁ x , func→ x₂ x
; ident = lift-eq (ident x₁) (ident x₂)
; distrib = lift-eq (distrib x₁) (distrib x₂)
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→: (identityNat F , .𝟙 {o = proj₂ c}) 𝔻 .𝟙
-- :ident: = begin
-- :func→: ((:obj: ×p ) .Product.obj .𝟙) ≡⟨⟩
-- :func→: (identityNat F , .𝟙) ≡⟨⟩
-- (identityTrans F C 𝔻⊕ F .func→ ( .𝟙)) ≡⟨⟩
-- (𝔻 .𝟙 𝔻⊕ F .func→ ( .𝟙)) ≡⟨ proj₂ 𝔻.ident ⟩
-- F .func→ ( .𝟙) ≡⟨ F .ident ⟩
-- 𝔻 .𝟙
-- where
-- _𝔻⊕_ = 𝔻 ._⊕_
-- open module 𝔻 = IsCategory (𝔻 .isCategory)
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 ._⊕_ {F×A} {G×B} {H×C}
module _
-- NaturalTransformation F G × .Arrow A B
{θ×α : NaturalTransformation F G × .Arrow A B}
{η×β : NaturalTransformation G H × .Arrow B C} where
θ : Transformation F G
θ = proj₁ (proj₁ θ×α)
θNat : Natural F G θ
θNat = proj₂ (proj₁ θ×α)
f : .Arrow A B
f = proj₂ θ×α
η : Transformation G H
η = proj₁ (proj₁ η×β)
ηNat : Natural G H η
ηNat = proj₂ (proj₁ η×β)
g : .Arrow B C
g = proj₂ η×β
-- :func→: ((θ , θNat) , f) = θB 𝔻⊕ F→f
_ : (:func→: {F×A} {G×B} θ×α) (θ B 𝔻⊕ F .func→ f)
_ = refl
ηθ : NaturalTransformation F H
ηθ = Fun ._⊕_ {F} {G} {H} (η , ηNat) (θ , θNat)
_ : ηθ Fun ._⊕_ {F} {G} {H} (η , ηNat) (θ , θNat)
_ = refl
ηθT = proj₁ ηθ
ηθN = proj₂ ηθ
_ : ηθT λ T η T 𝔻⊕ θ T -- Fun ._⊕_ {F} {G} {H} (η , ηNat) (θ , θNat)
_ = refl
:distrib: :
:func→: {F×A} {H×C} (η×β P⊕ θ×α)
(:func→: {G×B} {H×C} η×β) 𝔻⊕ (:func→: {F×A} {G×B} θ×α)
:distrib: = begin
:func→: {F×A} {H×C} (η×β P⊕ θ×α) ≡⟨⟩
:func→: {F×A} {H×C} (ηθ , g ℂ⊕ f) ≡⟨⟩
(ηθT C 𝔻⊕ F .func→ (g ℂ⊕ f)) ≡⟨ ηθN (g ℂ⊕ f)
(H .func→ (g ℂ⊕ f) 𝔻⊕ ηθT A) ≡⟨ cong (λ φ φ 𝔻⊕ ηθT A) (H .distrib)
((H .func→ g 𝔻⊕ H .func→ f) 𝔻⊕ ηθT A) ≡⟨ sym 𝔻.assoc
(H .func→ g 𝔻⊕ (H .func→ f 𝔻⊕ ηθT A)) ≡⟨⟩
(H .func→ g 𝔻⊕ (H .func→ f 𝔻⊕ (η A 𝔻⊕ θ 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)) ≡⟨⟩
((:func→: {G×B} {H×C} η×β) 𝔻⊕ (:func→: {F×A} {G×B} θ×α))
where
lemθ : θ B 𝔻⊕ F .func→ f G .func→ f 𝔻⊕ θ A
lemθ = θNat f
lemη : η C 𝔻⊕ G .func→ g H .func→ g 𝔻⊕ η B
lemη = ηNat g
lemm : ηθT C 𝔻⊕ F .func→ (g ℂ⊕ f) (H .func→ (g ℂ⊕ f) 𝔻⊕ ηθT A)
lemm = ηθN (g ℂ⊕ f)
final : η B 𝔻⊕ G .func→ f H .func→ f 𝔻⊕ η A
final = ηNat f
open module 𝔻 = IsCategory (𝔻 .isCategory)
-- Type of `:eval:` is aka.:
-- Functor ((:obj: ×p ) .Product.obj) 𝔻
-- :eval: : Cat .Arrow ((:obj: ×p ) .Product.obj) 𝔻
:eval: : Functor ((:obj: ×p ) .Product.obj) 𝔻
:eval: = record
{ func* = :func*:
; func→ = λ {dom} {cod} :func→: {dom} {cod}
; ident = λ {o} :ident: {o}
; distrib = λ {f u n k y} :distrib: {f} {u} {n} {k} {y}
}
-- Need to "lift equality of functors"
-- If I want to do this like I do it for pairs it's gonna be a pain.
postulate isUniqL : (Cat proj₁) x x₁
-- isUniqL = lift-eq-functors refl refl {!!} {!!}
module _ (𝔸 : Category ) (F : Functor ((𝔸 ×p ) .Product.obj) 𝔻) where
instance
CatHasProducts : HasProducts Cat
CatHasProducts = CatHasProducts {} {}
t : Cat .Arrow ((𝔸 ×p ) .Product.obj) 𝔻 Functor ((𝔸 ×p ) .Product.obj) 𝔻
t = refl
tt : Category
tt = (𝔸 ×p ) .Product.obj
open HasProducts CatHasProducts
postulate
transpose : Functor 𝔸 :obj:
eq : Cat ._⊕_ :eval: (parallelProduct transpose (Cat .𝟙 {o = })) F
postulate isUniqR : (Cat proj₂) x x₂
-- isUniqR = lift-eq-functors refl refl {!!} {!!}
catTranspose : ∃![ F~ ] (Cat ._⊕_ :eval: (parallelProduct F~ (Cat .𝟙 {o = })) F)
catTranspose = transpose , eq
isUniq : (Cat proj₁) x x₁ × (Cat proj₂) x x₂
isUniq = isUniqL , isUniqR
-- :isExponential: : IsExponential Cat A B :obj: {!:eval:!}
:isExponential: : IsExponential Cat 𝔻 :obj: :eval:
:isExponential: = catTranspose
uniq : ∃![ x ] ((Cat proj₁) x x₁ × (Cat proj₂) x x₂)
uniq = x , isUniq
-- :exponent: : Exponential (Cat ) A B
:exponent: : Exponential Cat 𝔻
:exponent: = record
{ obj = :obj:
; eval = :eval:
; isExponential = :isExponential:
}
instance
isProduct : IsProduct Cat proj₁ proj₂
isProduct = uniq
product : Product { = Cat} C D
product = record
{ obj = :product:
; proj₁ = proj₁
; proj₂ = proj₂
}
CatHasExponentials : HasExponentials Cat
CatHasExponentials = record { exponent = :exponent: }

View file

@ -13,7 +13,7 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
open Category
open Functor
module _ (F : Functor 𝔻) (G : Functor 𝔻) where
module _ (F G : Functor 𝔻) where
-- What do you call a non-natural tranformation?
Transformation : Set (c d')
Transformation = (C : .Object) 𝔻 .Arrow (F .func* C) (G .func* C)

View file

@ -11,16 +11,35 @@ open import Cat.Category
open import Cat.Functor
open Category
Sets : { : Level} Category (lsuc )
Sets {} = record
{ Object = Set
; Arrow = λ T U T U
; 𝟙 = id
; _⊕_ = _∘_
; isCategory = record { assoc = refl ; ident = funExt (λ _ refl) , funExt (λ _ refl) }
}
where
open import Function
module _ { : Level} where
Sets : Category (lsuc )
Sets = record
{ Object = Set
; Arrow = λ T U T U
; 𝟙 = id
; _⊕_ = _∘_
; isCategory = record { assoc = refl ; ident = funExt (λ _ refl) , funExt (λ _ refl) }
}
where
open import Function
private
module _ {X A B : Set } (f : X A) (g : X B) where
pair : (X A × B)
pair x = f x , g x
lem : Sets ._⊕_ proj₁ pair f × Sets ._⊕_ snd pair g
proj₁ lem = refl
snd lem = refl
instance
isProduct : {A B : Sets .Object} IsProduct Sets {A} {B} fst snd
isProduct f g = pair f g , lem f g
product : (A B : Sets .Object) Product { = Sets} A B
product A B = record { obj = A × B ; proj₁ = fst ; proj₂ = snd ; isProduct = {!!} }
instance
SetsHasProducts : HasProducts Sets
SetsHasProducts = record { product = product }
-- Covariant Presheaf
Representable : { ' : Level} ( : Category ') Set ( lsuc ')

View file

@ -88,6 +88,7 @@ record Product { ' : Level} { : Category '} (A B : .Object)
proj₁ : .Arrow obj A
proj₂ : .Arrow obj B
{{isProduct}} : IsProduct proj₁ proj₂
arrowProduct : {X} (π₁ : Arrow X A) (π₂ : Arrow X B)
Arrow X obj
arrowProduct π₁ π₂ = fst (isProduct π₁ π₂)
@ -96,6 +97,18 @@ record HasProducts { ' : Level} ( : Category ') : Set (
field
product : (A B : .Object) Product { = } A B
open Product
objectProduct : (A B : .Object) .Object
objectProduct A B = Product.obj (product A B)
-- The product mentioned in awodey in Def 6.1 is not the regular product of arrows.
-- It's a "parallel" product
parallelProduct : {A A' B B' : .Object} .Arrow A A' .Arrow B B'
.Arrow (objectProduct A B) (objectProduct A' B')
parallelProduct {A = A} {A' = A'} {B = B} {B' = B'} a b = arrowProduct (product A' B')
( ._⊕_ a ((product A B) .proj₁))
( ._⊕_ b ((product A B) .proj₂))
module _ { ' : Level} ( : Category ') where
Opposite : Category '
Opposite =
@ -127,3 +140,37 @@ module _ { ' : Level} { : Category '} where
HomFromArrow : (A : .Object) {B B' : .Object} (g : .Arrow B B')
Hom A B Hom A B'
HomFromArrow _A = _⊕_
module _ { '} ( : Category ') {{HasProducts : HasProducts }} where
open HasProducts HasProducts
open Product hiding (obj)
private
_×p_ : (A B : .Object) .Object
_×p_ A B = Product.obj (product A B)
module _ (B C : .Category.Object) where
IsExponential : (Cᴮ : .Object) .Arrow (Cᴮ ×p B) C Set ( ')
IsExponential Cᴮ eval = (A : .Object) (f : .Arrow (A ×p B) C)
∃![ f~ ] ( ._⊕_ eval (parallelProduct f~ ( .𝟙)) f)
record Exponential : Set ( ') where
field
-- obj ≡ Cᴮ
obj : .Object
eval : .Arrow ( obj ×p B ) C
{{isExponential}} : IsExponential obj eval
-- If I make this an instance-argument then the instance resolution
-- algorithm goes into an infinite loop. Why?
productsFromExp : HasProducts
productsFromExp = HasProducts
transpose : (A : .Object) .Arrow (A ×p B) C .Arrow A obj
transpose A f = fst (isExponential A f)
record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where
field
exponent : (A B : .Object) Exponential A B
record CartesianClosed { ' : Level} ( : Category ') : Set ( ') where
field
{{hasProducts}} : HasProducts
{{hasExponentials}} : HasExponentials

View file

@ -1,4 +1,4 @@
{-# OPTIONS --allow-unsolved-metas #-}
{-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Category.Properties where
@ -48,33 +48,25 @@ epi-mono-is-not-iso f =
in {!!}
-}
open import Cat.Categories.Cat
open Exponential
open HasExponentials CatHasExponentials
module _ { '} ( : Category ') {{hasProducts : HasProducts }} (B C : .Category.Object) where
open Category
open HasProducts hasProducts
open Product
prod-obj : (A B : .Object) .Object
prod-obj A B = Product.obj (product A B)
-- The product mentioned in awodey in Def 6.1 is not the regular product of arrows.
-- It's a "parallel" product
×A : {A A' B B' : .Object} .Arrow A A' .Arrow B B'
.Arrow (prod-obj A B) (prod-obj A' B')
×A {A = A} {A' = A'} {B = B} {B' = B'} a b = arrowProduct (product A' B')
( ._⊕_ a ((product A B) .proj₁))
( ._⊕_ b ((product A B) .proj₂))
Exp : Set {!!}
Exp = Exponential (Cat {!!} {!!}) {{HasProducts = {!!}}}
Sets (Opposite {!!})
IsExponential : {Cᴮ : .Object} .Arrow (prod-obj Cᴮ B) C Set ( ')
IsExponential eval = (A : .Object) (f : .Arrow (prod-obj A B) C)
∃![ f~ ] ( ._⊕_ eval (×A f~ ( .𝟙)) f)
-- _⇑_ : (A B : Cat .Object) → Cat .Object
-- A ⇑ B = (exponent A B) .obj
record Exponential : Set ( ') where
field
-- obj ≡ Cᴮ
obj : .Object
eval : .Arrow ( prod-obj obj B ) C
{{isExponential}} : IsExponential eval
-- private
-- :func*: : .Object (Sets Opposite ) .Object
-- :func*: x = {!!}
_⇑_ = Exponential
-- yoneda : ∀ { '} → { : Category '} → Functor (Sets ⇑ (Opposite ))
-- yoneda = {!!}
-- yoneda : Functor (Sets ⇑ (Opposite ))
-- yoneda = record
-- { func* = :func*:
-- ; func→ = {!!}
-- ; ident = {!!}
-- ; distrib = {!!}
-- }