Changes to the category of categories

This commit is contained in:
Frederik Hanghøj Iversen 2018-02-05 16:35:33 +01:00
parent e8215b2c05
commit e8ac6786ff
5 changed files with 154 additions and 111 deletions

View file

@ -9,7 +9,9 @@ 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.Category.Functor
open import Cat.Category.Product
open import Cat.Category.Exponential
open import Cat.Equality
open Equality.Data.Product
@ -26,12 +28,12 @@ module _ ( ' : Level) where
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 ])
(λ 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})
: (λ i {A : Object 𝔸} eq→ i (𝟙 𝔸 {A}) 𝟙 𝔻 {eq* i A})
[ (H ∘f (G ∘f F)) .isFunctor .ident
((H ∘f G) ∘f F) .isFunctor .ident
]
@ -58,12 +60,12 @@ module _ ( ' : Level) where
eq→ = refl
postulate
eqI-r
: (λ i {c : .Object} (λ _ 𝔻 [ func* F c , func* F c ])
[ func→ F ( .𝟙) 𝔻 .𝟙 ])
: (λ 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}
{A B C : Object } {f : [ A , B ]} {g : [ 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
@ -73,40 +75,50 @@ module _ ( ' : Level) where
postulate
eq* : (identity ∘f F) .func* F .func*
eq→ : PathP
(λ i {x y : Object } .Arrow x y 𝔻 .Arrow (eq* i x) (eq* i y))
(λ i {x y : Object } [ x , y ] 𝔻 [ eq* i x , eq* i y ])
((identity ∘f F) .func→) (F .func→)
eqI : (λ i {A : .Object} eq→ i ( .𝟙 {A}) 𝔻 .𝟙 {eq* i A})
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}
eqD : PathP (λ i {A B C : Object } {f : [ A , B ]} {g : [ 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
RawCat : RawCategory (lsuc ( ')) ( ')
RawCat =
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
-- }
}
}
open IsCategory
instance
:isCategory: : IsCategory RawCat
assoc :isCategory: {f = F} {G} {H} = assc {F = F} {G = G} {H = H}
ident :isCategory: = ident-r , ident-l
arrow-is-set :isCategory: = {!!}
univalent :isCategory: = {!!}
Cat : Category (lsuc ( ')) ( ')
raw Cat = RawCat
module _ { ' : Level} where
module _ ( 𝔻 : Category ') where
private
Catt = Cat '
:Object: = .Object × 𝔻 .Object
: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
@ -114,25 +126,35 @@ module _ { ' : Level} where
:Arrow: a c
_:⊕:_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) [ bc∈C ab∈C ] , 𝔻 [ bc∈D ab∈D ]}
:rawProduct: : RawCategory '
RawCategory.Object :rawProduct: = :Object:
RawCategory.Arrow :rawProduct: = :Arrow:
RawCategory.𝟙 :rawProduct: = :𝟙:
RawCategory._∘_ :rawProduct: = _:⊕:_
module C = IsCategory ( .isCategory)
module D = IsCategory (𝔻 .isCategory)
postulate
issSet : {A B : RawCategory.Object :rawProduct:} isSet (RawCategory.Arrow :rawProduct: A B)
instance
:isCategory: : IsCategory :Object: :Arrow: :𝟙: _:⊕:_
:isCategory: = record
{ assoc = Σ≡ C.assoc D.assoc
; ident
:isCategory: : IsCategory :rawProduct:
-- :isCategory: = record
-- { assoc = Σ≡ C.assoc D.assoc
-- ; ident
-- = Σ≡ (fst C.ident) (fst D.ident)
-- , Σ≡ (snd C.ident) (snd D.ident)
-- ; arrow-is-set = issSet
-- ; univalent = {!!}
-- }
IsCategory.assoc :isCategory: = Σ≡ C.assoc D.assoc
IsCategory.ident :isCategory:
= Σ≡ (fst C.ident) (fst D.ident)
, Σ≡ (snd C.ident) (snd D.ident)
}
where
open module C = IsCategory ( .isCategory)
open module D = IsCategory (𝔻 .isCategory)
IsCategory.arrow-is-set :isCategory: = issSet
IsCategory.univalent :isCategory: = {!!}
:product: : Category '
:product: = record
{ Object = :Object:
; Arrow = :Arrow:
; 𝟙 = :𝟙:
; _∘_ = _:⊕:_
}
raw :product: = :rawProduct:
proj₁ : Catt [ :product: , ]
proj₁ = record { func* = fst ; func→ = fst ; isFunctor = record { ident = refl ; distrib = refl } }
@ -143,28 +165,32 @@ module _ { ' : Level} where
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)
postulate 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
-- Turned into postulate after:
-- > commit e8215b2c051062c6301abc9b3f6ec67106259758 (HEAD -> dev, github/dev)
-- > Author: Frederik Hanghøj Iversen <fhi.1990@gmail.com>
-- > Date: Mon Feb 5 14:59:53 2018 +0100
postulate isUniqL : Catt [ proj₁ x ] x₁
-- isUniqL = Functor≡ eq* eq→ {!!}
-- where
-- eq* : (Catt [ proj₁ ∘ x ]) .func* ≡ x₁ .func*
-- eq* = {!refl!}
-- eq→ : (λ i → {A : Object X} {B : Object X} → 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₂
@ -202,55 +228,55 @@ module _ ( : Level) where
Cat = Cat
module _ ( 𝔻 : Category ) where
private
:obj: : Cat .Object
:obj: : Object (Cat )
:obj: = Fun { = } {𝔻 = 𝔻}
:func*: : Functor 𝔻 × .Object 𝔻 .Object
:func*: : Functor 𝔻 × Object Object 𝔻
:func*: (F , A) = F .func* A
module _ {dom cod : Functor 𝔻 × .Object} where
module _ {dom cod : Functor 𝔻 × Object } where
private
F : Functor 𝔻
F = proj₁ dom
A : .Object
A : Object
A = proj₂ dom
G : Functor 𝔻
G = proj₁ cod
B : .Object
B : Object
B = proj₂ cod
:func→: : (pobj : NaturalTransformation F G × .Arrow A B)
𝔻 .Arrow (F .func* A) (G .func* B)
:func→: : (pobj : NaturalTransformation F G × [ A , B ])
𝔻 [ F .func* A , G .func* B ]
:func→: ((θ , θNat) , f) = result
where
θA : 𝔻 .Arrow (F .func* A) (G .func* A)
θA : 𝔻 [ F .func* A , G .func* A ]
θA = θ A
θB : 𝔻 .Arrow (F .func* B) (G .func* B)
θB : 𝔻 [ F .func* B , G .func* B ]
θB = θ B
F→f : 𝔻 .Arrow (F .func* A) (F .func* B)
F→f : 𝔻 [ F .func* A , F .func* B ]
F→f = F .func→ f
G→f : 𝔻 .Arrow (G .func* A) (G .func* B)
G→f : 𝔻 [ G .func* A , G .func* B ]
G→f = G .func→ f
l : 𝔻 .Arrow (F .func* A) (G .func* B)
l : 𝔻 [ F .func* A , G .func* B ]
l = 𝔻 [ θB F→f ]
r : 𝔻 .Arrow (F .func* A) (G .func* B)
r : 𝔻 [ 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 : 𝔻 [ F .func* A , G .func* B ]
result = l
_×p_ = product
module _ {c : Functor 𝔻 × .Object} where
module _ {c : Functor 𝔻 × Object } where
private
F : Functor 𝔻
F = proj₁ c
C : .Object
C : Object
C = proj₂ c
-- NaturalTransformation F G × .Arrow A B
@ -259,19 +285,19 @@ module _ ( : Level) where
-- where
-- open module 𝔻 = IsCategory (𝔻 .isCategory)
-- Unfortunately the equational version has some ambigous arguments.
:ident: : :func→: {c} {c} (identityNat F , .𝟙 {o = proj₂ c}) 𝔻 .𝟙
: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
𝔻 .𝟙
:func→: {c} {c} (𝟙 (Product.obj (:obj: ×p )) {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
module _ {F×A G×B H×C : Functor 𝔻 × Object } where
F = F×A .proj₁
A = F×A .proj₂
G = G×B .proj₁
@ -279,27 +305,27 @@ module _ ( : Level) where
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}
_P⊕_ = Category._∘_ (Product.obj (:obj: ×p )) {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
{θ×f : NaturalTransformation F G × [ A , B ]}
{η×g : NaturalTransformation G H × [ B , C ]} where
private
θ : Transformation F G
θ = proj₁ (proj₁ θ×f)
θNat : Natural F G θ
θNat = proj₂ (proj₁ θ×f)
f : .Arrow A B
f : [ A , B ]
f = proj₂ θ×f
η : Transformation G H
η = proj₁ (proj₁ η×g)
ηNat : Natural G H η
ηNat = proj₂ (proj₁ η×g)
g : .Arrow B C
g : [ B , C ]
g = proj₂ η×g
ηθNT : NaturalTransformation F H
ηθNT = Fun .Category._∘_ {F} {G} {H} (η , ηNat) (θ , θNat)
ηθNT = Category._∘_ Fun {F} {G} {H} (η , ηNat) (θ , θNat)
ηθ = proj₁ ηθNT
ηθNat = proj₂ ηθNT
@ -341,17 +367,28 @@ module _ ( : Level) where
}
module _ (𝔸 : Category ) (F : Functor ((𝔸 ×p ) .Product.obj) 𝔻) where
open HasProducts (hasProducts {} {}) using (parallelProduct)
open HasProducts (hasProducts {} {}) renaming (_|×|_ to parallelProduct)
postulate
transpose : Functor 𝔸 :obj:
eq : Cat [ :eval: (parallelProduct transpose (Cat .𝟙 {o = })) ] F
eq : Cat [ :eval: (parallelProduct transpose (𝟙 Cat {o = })) ] F
-- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (𝟙 Cat {o = })) ] ≡ F
-- eq' : (Cat [ :eval: ∘
-- (record { product = product } HasProducts.|×| transpose)
-- (𝟙 Cat)
-- ])
-- ≡ F
catTranspose : ∃![ F~ ] (Cat [ :eval: (parallelProduct F~ (Cat .𝟙 {o = }))] F )
catTranspose = transpose , eq
-- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758`
-- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Cat [
-- :eval: (parallelProduct F~ (𝟙 Cat {o = }))] F) catTranspose =
-- transpose , eq
:isExponential: : IsExponential Cat 𝔻 :obj: :eval:
:isExponential: = catTranspose
:isExponential: = {!catTranspose!}
where
open HasProducts (hasProducts {} {}) using (_|×|_)
-- :isExponential: = λ 𝔸 F transpose 𝔸 F , eq' 𝔸 F
-- :exponent: : Exponential (Cat ) A B
:exponent: : Exponential Cat 𝔻

View file

@ -76,8 +76,10 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
-- TODO: might want to implement isEquiv differently, there are 3
-- equivalent formulations in the book.
Univalent : Set (a b)
Univalent = {A B : Object} isEquiv (A B) (A B) (id-to-iso A B)
field
univalent : {A B : Object} isEquiv (A B) (A B) (id-to-iso A B)
univalent : Univalent
module _ {A B : Object} where
Epimorphism : {X : Object } (f : Arrow A B) Set b

View file

@ -19,7 +19,7 @@ module _ { '} ( : Category ') {{hasProducts : HasProducts }}
module _ (B C : Object ) where
IsExponential : (Cᴮ : Object ) [ Cᴮ ×p B , C ] Set ( ')
IsExponential Cᴮ eval = (A : Object ) (f : [ A ×p B , C ])
∃![ f~ ] ( [ eval parallelProduct f~ (Category.𝟙 )] f)
∃![ f~ ] ( [ eval f~ |×| Category.𝟙 ] f)
record Exponential : Set ( ') where
field

View file

@ -1,8 +1,8 @@
module Cat.Category.Product where
open import Agda.Primitive
open import Data.Product
open import Cubical
open import Data.Product as P hiding (_×_)
open import Cat.Category
@ -12,14 +12,16 @@ module _ { ' : Level} ( : Category ') {A B obj : Object } whe
IsProduct : (π₁ : [ obj , A ]) (π₂ : [ obj , B ]) Set ( ')
IsProduct π₁ π₂
= {X : Object } (x₁ : [ X , A ]) (x₂ : [ X , B ])
∃![ x ] ( [ π₁ x ] x₁ × [ π₂ x ] x₂)
∃![ x ] ( [ π₁ x ] x₁ P.× [ π₂ x ] x₂)
-- Tip from Andrea; Consider this style for efficiency:
-- record IsProduct { ' : Level} ( : Category {} {'})
-- {A B obj : Object } (π₁ : Arrow obj A) (π₂ : Arrow obj B) : Set (') where
-- record IsProduct {a b : Level} ( : Category a b)
-- {A B obj : Object } (π₁ : Arrow obj A) (π₂ : Arrow obj B) : Set (a ⊔ b) where
-- field
-- isProduct : ∀ {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B)
-- → ∃![ x ] ( ._⊕_ π₁ x ≡ x₁ × . _⊕_ π₂ x ≡ x₂)
-- issProduct : ∀ {X : Object } (x₁ : [ X , A ]) (x₂ : [ X , B ])
-- → ∃![ x ] ( [ π₁ ∘ x ] ≡ x₁ P.× [ π₂ ∘ x ] ≡ x₂)
-- open IsProduct
record Product { ' : Level} { : Category '} (A B : Object ) : Set ( ') where
no-eta-equality
@ -29,9 +31,9 @@ record Product { ' : Level} { : Category '} (A B : Object ) :
proj₂ : [ obj , B ]
{{isProduct}} : IsProduct proj₁ proj₂
arrowProduct : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
_P[_×_] : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
[ X , obj ]
arrowProduct π₁ π₂ = proj₁ (isProduct π₁ π₂)
_P[_×_] π₁ π₂ = proj₁ (isProduct π₁ π₂)
record HasProducts { ' : Level} ( : Category ') : Set ( ') where
field
@ -39,12 +41,14 @@ record HasProducts { ' : Level} ( : Category ') : Set (
open Product
objectProduct : (A B : Object ) Object
objectProduct A B = Product.obj (product A B)
_×_ : (A B : Object ) Object
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 } [ A , A' ] [ B , B' ]
[ 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₂ ])
_|×|_ : {A A' B B' : Object } [ A , A' ] [ B , B' ]
[ A × B , A' × B' ]
_|×|_ {A = A} {A' = A'} {B = B} {B' = B'} a b
= product A' B'
P[ [ a (product A B) .proj₁ ]
× [ b (product A B) .proj₂ ]
]

View file

@ -60,7 +60,7 @@ open Functor
-- open import Cat.Categories.Fun
-- open import Cat.Categories.Sets
-- -- module Cat = Cat.Categories.Cat
-- open Exponential
-- open import Cat.Category.Exponential
-- private
-- Cat = Cat
-- prshf = presheaf { = }