Merge branch 'dev'

This commit is contained in:
Frederik Hanghøj Iversen 2018-02-06 14:31:28 +01:00
commit 4df4231906
17 changed files with 615 additions and 441 deletions

View file

@ -1,14 +1,19 @@
module Cat where module Cat where
import Cat.Category import Cat.Category
import Cat.Functor
import Cat.CwF import Cat.CwF
import Cat.Category.Functor
import Cat.Category.Product
import Cat.Category.Exponential
import Cat.Category.CartesianClosed
import Cat.Category.Pathy import Cat.Category.Pathy
import Cat.Category.Bij import Cat.Category.Bij
import Cat.Category.Free
import Cat.Category.Properties import Cat.Category.Properties
import Cat.Categories.Sets import Cat.Categories.Sets
-- import Cat.Categories.Cat -- import Cat.Categories.Cat
import Cat.Categories.Rel import Cat.Categories.Rel
import Cat.Categories.Free
import Cat.Categories.Fun import Cat.Categories.Fun
import Cat.Categories.Cube import Cat.Categories.Cube

View file

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

View file

@ -13,7 +13,7 @@ open import Relation.Nullary
open import Relation.Nullary.Decidable open import Relation.Nullary.Decidable
open import Cat.Category open import Cat.Category
open import Cat.Functor open import Cat.Category.Functor
open import Cat.Equality open import Cat.Equality
open Equality.Data.Product open Equality.Data.Product
@ -65,15 +65,16 @@ module _ { ' : Level} (Ns : Set ) where
Hom = Σ Hom' rules Hom = Σ Hom' rules
module Raw = RawCategory
-- The category of names and substitutions -- The category of names and substitutions
Raw : RawCategory -- o (lsuc lzero ⊔ o) Raw : RawCategory -- o (lsuc lzero ⊔ o)
Raw = record Raw.Object Raw = FiniteDecidableSubset
{ Object = FiniteDecidableSubset Raw.Arrow Raw = Hom
-- { Object = Ns → Bool Raw.𝟙 Raw {o} = inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} }
; Arrow = Hom Raw._∘_ Raw = {!!}
; 𝟙 = λ { {o} inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} } }
; _∘_ = {!!} postulate IsCategory : IsCategory Raw
}
postulate RawIsCategory : IsCategory Raw
: Category : Category
= Raw , RawIsCategory raw = Raw
isCategory = IsCategory

View file

@ -46,9 +46,9 @@ module _ (a b : Level) where
isCategory = record isCategory = record
{ assoc = λ {A} {B} {C} {D} {f} {g} {h} assoc {D = D} {f} {g} {h} { assoc = λ {A} {B} {C} {D} {f} {g} {h} assoc {D = D} {f} {g} {h}
; ident = λ {A} {B} {f} ident {A} {B} {f = f} ; ident = λ {A} {B} {f} ident {A} {B} {f = f}
; arrow-is-set = ? ; arrowIsSet = {!!}
; univalent = ? ; univalent = {!!}
} }
Fam : Category (lsuc (a b)) (a b) Fam : Category (lsuc (a b)) (a b)
Fam = RawFam , isCategory Category.raw Fam = RawFam

View file

@ -0,0 +1,62 @@
{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Categories.Free where
open import Agda.Primitive
open import Cubical hiding (Path ; isSet ; empty)
open import Data.Product
open import Cat.Category
open IsCategory
open Category
-- data Path { : Level} {A : Set } : (a b : A) → Set where
-- emptyPath : {a : A} → Path a a
-- concatenate : {a b c : A} → Path a b → Path b c → Path a b
module _ { ' : Level} ( : Category ') where
module = Category
-- import Data.List
-- P : (a b : Object ) → Set (')
-- P = {!Data.List.List ?!}
-- Generalized paths:
-- data P { : Level} {A : Set } (R : A → A → Set ) : (a b : A) → Set where
-- e : {a : A} → P R a a
-- c : {a b c : A} → R a b → P R b c → P R a c
-- Path's are like lists with directions.
-- This implementation is specialized to categories.
data Path : (a b : Object ) Set ( ') where
empty : {A : Object } Path A A
cons : {A B C} [ B , C ] Path A B Path A C
concatenate : {A B C : Object } Path B C Path A B Path A C
concatenate empty p = p
concatenate (cons x q) p = cons x (concatenate q p)
private
module _ {A B C D : Object } where
p-assoc : {r : Path A B} {q : Path B C} {p : Path C D} concatenate p (concatenate q r) concatenate (concatenate p q) r
p-assoc {r} {q} {p} = {!!}
module _ {A B : Object } {p : Path A B} where
-- postulate
-- ident-r : concatenate {A} {A} {B} p (lift 𝟙) ≡ p
-- ident-l : concatenate {A} {B} {B} (lift 𝟙) p ≡ p
module _ {A B : Object } where
isSet : IsSet (Path A B)
isSet = {!!}
RawFree : RawCategory ( ')
RawFree = record
{ Object = Object
; Arrow = Path
; 𝟙 = λ {o} {!lift 𝟙!}
; _∘_ = λ {a b c} {!concatenate {a} {b} {c}!}
}
RawIsCategoryFree : IsCategory RawFree
RawIsCategoryFree = record
{ assoc = {!p-assoc!}
; ident = {!ident-r , ident-l!}
; arrowIsSet = {!!}
; univalent = {!!}
}

View file

@ -7,22 +7,25 @@ open import Function
open import Data.Product open import Data.Product
open import Cat.Category open import Cat.Category
open import Cat.Functor open import Cat.Category.Functor
module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Category d d'} where module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Category d d'} where
open Category hiding ( _∘_ ; Arrow ) open Category hiding ( _∘_ ; Arrow )
open Functor open Functor
module _ (F G : Functor 𝔻) where module _ (F G : Functor 𝔻) where
private
module F = Functor F
module G = Functor G
-- What do you call a non-natural tranformation? -- What do you call a non-natural tranformation?
Transformation : Set (c d') Transformation : Set (c d')
Transformation = (C : .Object) 𝔻 [ F .func* C , G .func* C ] Transformation = (C : Object ) 𝔻 [ F.func* C , G.func* C ]
Natural : Transformation Set (c (c' d')) Natural : Transformation Set (c (c' d'))
Natural θ Natural θ
= {A B : .Object} = {A B : Object }
(f : [ A , B ]) (f : [ A , B ])
𝔻 [ θ B F .func→ f ] 𝔻 [ G .func→ f θ A ] 𝔻 [ θ B F.func→ f ] 𝔻 [ G.func→ f θ A ]
NaturalTransformation : Set (c c' d') NaturalTransformation : Set (c c' d')
NaturalTransformation = Σ Transformation Natural NaturalTransformation = Σ Transformation Natural
@ -30,29 +33,29 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
-- NaturalTranformation : Set (c ⊔ (c' ⊔ d')) -- NaturalTranformation : Set (c ⊔ (c' ⊔ d'))
-- NaturalTranformation = ∀ (θ : Transformation) {A B : .Object} → (f : .Arrow A B) → 𝔻 ._⊕_ (θ B) (F .func→ f) ≡ 𝔻 ._⊕_ (G .func→ f) (θ A) -- NaturalTranformation = ∀ (θ : Transformation) {A B : .Object} → (f : .Arrow A B) → 𝔻 ._⊕_ (θ B) (F .func→ f) ≡ 𝔻 ._⊕_ (G .func→ f) (θ A)
module _ {F G : Functor 𝔻} where NaturalTransformation≡ : {α β : NaturalTransformation}
NaturalTransformation≡ : {α β : NaturalTransformation F G}
(eq₁ : α .proj₁ β .proj₁) (eq₁ : α .proj₁ β .proj₁)
(eq₂ : PathP (eq₂ : PathP
(λ i {A B : .Object} (f : [ A , B ]) (λ i {A B : Object } (f : [ A , B ])
𝔻 [ eq₁ i B F .func→ f ] 𝔻 [ eq₁ i B F.func→ f ]
𝔻 [ G .func→ f eq₁ i A ]) 𝔻 [ G.func→ f eq₁ i A ])
(α .proj₂) (β .proj₂)) (α .proj₂) (β .proj₂))
α β α β
NaturalTransformation≡ eq₁ eq₂ i = eq₁ i , eq₂ i NaturalTransformation≡ eq₁ eq₂ i = eq₁ i , eq₂ i
identityTrans : (F : Functor 𝔻) Transformation F F identityTrans : (F : Functor 𝔻) Transformation F F
identityTrans F C = 𝔻 .𝟙 identityTrans F C = 𝟙 𝔻
identityNatural : (F : Functor 𝔻) Natural F F (identityTrans F) identityNatural : (F : Functor 𝔻) Natural F F (identityTrans F)
identityNatural F {A = A} {B = B} f = begin identityNatural F {A = A} {B = B} f = begin
𝔻 [ identityTrans F B F→ f ] ≡⟨⟩ 𝔻 [ identityTrans F B F→ f ] ≡⟨⟩
𝔻 [ 𝔻 .𝟙 F→ f ] ≡⟨ proj₂ 𝔻.ident 𝔻 [ 𝟙 𝔻 F→ f ] ≡⟨ proj₂ 𝔻.ident
F→ f ≡⟨ sym (proj₁ 𝔻.ident) F→ f ≡⟨ sym (proj₁ 𝔻.ident)
𝔻 [ F→ f 𝔻 .𝟙 ] ≡⟨⟩ 𝔻 [ F→ f 𝟙 𝔻 ] ≡⟨⟩
𝔻 [ F→ f identityTrans F A ] 𝔻 [ F→ f identityTrans F A ]
where where
F→ = F .func→ module F = Functor F
F→ = F.func→
module 𝔻 = IsCategory (isCategory 𝔻) module 𝔻 = IsCategory (isCategory 𝔻)
identityNat : (F : Functor 𝔻) NaturalTransformation F F identityNat : (F : Functor 𝔻) NaturalTransformation F F
@ -60,20 +63,23 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
module _ {F G H : Functor 𝔻} where module _ {F G H : Functor 𝔻} where
private private
module F = Functor F
module G = Functor G
module H = Functor H
_∘nt_ : Transformation G H Transformation F G Transformation F H _∘nt_ : Transformation G H Transformation F G Transformation F H
(θ ∘nt η) C = 𝔻 [ θ C η C ] (θ ∘nt η) C = 𝔻 [ θ C η C ]
NatComp _:⊕:_ : NaturalTransformation G H NaturalTransformation F G NaturalTransformation F H NatComp _:⊕:_ : NaturalTransformation G H NaturalTransformation F G NaturalTransformation F H
proj₁ ((θ , _) :⊕: (η , _)) = θ ∘nt η proj₁ ((θ , _) :⊕: (η , _)) = θ ∘nt η
proj₂ ((θ , θNat) :⊕: (η , ηNat)) {A} {B} f = begin proj₂ ((θ , θNat) :⊕: (η , ηNat)) {A} {B} f = begin
𝔻 [ (θ ∘nt η) B F .func→ f ] ≡⟨⟩ 𝔻 [ (θ ∘nt η) B F.func→ f ] ≡⟨⟩
𝔻 [ 𝔻 [ θ B η B ] F .func→ f ] ≡⟨ sym assoc 𝔻 [ 𝔻 [ θ B η B ] F.func→ f ] ≡⟨ sym assoc
𝔻 [ θ B 𝔻 [ η B F .func→ f ] ] ≡⟨ cong (λ φ 𝔻 [ θ B φ ]) (ηNat f) 𝔻 [ θ B 𝔻 [ η B F.func→ f ] ] ≡⟨ cong (λ φ 𝔻 [ θ B φ ]) (ηNat f)
𝔻 [ θ B 𝔻 [ G .func→ f η A ] ] ≡⟨ assoc 𝔻 [ θ B 𝔻 [ G.func→ f η A ] ] ≡⟨ assoc
𝔻 [ 𝔻 [ θ B G .func→ f ] η A ] ≡⟨ cong (λ φ 𝔻 [ φ η A ]) (θNat f) 𝔻 [ 𝔻 [ θ B G.func→ f ] η A ] ≡⟨ cong (λ φ 𝔻 [ φ η A ]) (θNat f)
𝔻 [ 𝔻 [ H .func→ f θ A ] η A ] ≡⟨ sym assoc 𝔻 [ 𝔻 [ H.func→ f θ A ] η A ] ≡⟨ sym assoc
𝔻 [ H .func→ f 𝔻 [ θ A η A ] ] ≡⟨⟩ 𝔻 [ H.func→ f 𝔻 [ θ A η A ] ] ≡⟨⟩
𝔻 [ H .func→ f (θ ∘nt η) A ] 𝔻 [ H.func→ f (θ ∘nt η) A ]
where where
open IsCategory (isCategory 𝔻) open IsCategory (isCategory 𝔻)
@ -110,12 +116,12 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
:isCategory: = record :isCategory: = record
{ assoc = λ {A B C D} :assoc: {A} {B} {C} {D} { assoc = λ {A B C D} :assoc: {A} {B} {C} {D}
; ident = λ {A B} :ident: {A} {B} ; ident = λ {A B} :ident: {A} {B}
; arrow-is-set = ? ; arrowIsSet = {!!}
; univalent = ? ; univalent = {!!}
} }
Fun : Category (c c' d d') (c c' d') Fun : Category (c c' d d') (c c' d')
Fun = RawFun , :isCategory: raw Fun = RawFun
module _ { ' : Level} ( : Category ') where module _ { ' : Level} ( : Category ') where
open import Cat.Categories.Sets open import Cat.Categories.Sets

View file

@ -166,6 +166,6 @@ RawIsCategoryRel : IsCategory RawRel
RawIsCategoryRel = record RawIsCategoryRel = record
{ assoc = funExt is-assoc { assoc = funExt is-assoc
; ident = funExt ident-l , funExt ident-r ; ident = funExt ident-l , funExt ident-r
; arrow-is-set = {!!} ; arrowIsSet = {!!}
; univalent = {!!} ; univalent = {!!}
} }

View file

@ -1,4 +1,4 @@
{-# OPTIONS --allow-unsolved-metas #-} {-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Categories.Sets where module Cat.Categories.Sets where
open import Cubical open import Cubical
@ -7,28 +7,28 @@ open import Data.Product
import Function import Function
open import Cat.Category open import Cat.Category
open import Cat.Functor open import Cat.Category.Functor
open import Cat.Category.Product
open Category open Category
module _ { : Level} where module _ { : Level} where
SetsRaw : RawCategory (lsuc ) SetsRaw : RawCategory (lsuc )
SetsRaw = record RawCategory.Object SetsRaw = Set
{ Object = Set RawCategory.Arrow SetsRaw = λ T U T U
; Arrow = λ T U T U RawCategory.𝟙 SetsRaw = Function.id
; 𝟙 = Function.id RawCategory._∘_ SetsRaw = Function._∘_
; _∘_ = Function._∘_
}
open IsCategory
SetsIsCategory : IsCategory SetsRaw SetsIsCategory : IsCategory SetsRaw
SetsIsCategory = record assoc SetsIsCategory = refl
{ assoc = refl proj₁ (ident SetsIsCategory) = funExt λ _ refl
; ident = funExt (λ _ refl) , funExt (λ _ refl) proj₂ (ident SetsIsCategory) = funExt λ _ refl
; arrow-is-set = {!!} arrowIsSet SetsIsCategory = {!!}
; univalent = {!!} univalent SetsIsCategory = {!!}
}
Sets : Category (lsuc ) Sets : Category (lsuc )
Sets = SetsRaw , SetsIsCategory raw Sets = SetsRaw
isCategory Sets = SetsIsCategory
private private
module _ {X A B : Set } (f : X A) (g : X B) where module _ {X A B : Set } (f : X A) (g : X B) where
@ -39,10 +39,10 @@ module _ { : Level} where
proj₁ lem = refl proj₁ lem = refl
proj₂ lem = refl proj₂ lem = refl
instance instance
isProduct : {A B : Sets .Object} IsProduct Sets {A} {B} proj₁ proj₂ isProduct : {A B : Object Sets} IsProduct Sets {A} {B} proj₁ proj₂
isProduct f g = f &&& g , lem f g isProduct f g = f &&& g , lem f g
product : (A B : Sets .Object) Product { = Sets} A B product : (A B : Object Sets) Product { = Sets} A B
product A B = record { obj = A × B ; proj₁ = proj₁ ; proj₂ = proj₂ ; isProduct = isProduct } product A B = record { obj = A × B ; proj₁ = proj₁ ; proj₂ = proj₂ ; isProduct = isProduct }
instance instance
@ -56,8 +56,10 @@ Representable {' = '} = Functor (Sets {'})
-- The "co-yoneda" embedding. -- The "co-yoneda" embedding.
representable : { '} { : Category '} Category.Object Representable representable : { '} { : Category '} Category.Object Representable
representable { = } A = record representable { = } A = record
{ func* = λ B .Arrow A B { raw = record
; func→ = ._∘_ { func* = λ B [ A , B ]
; func→ = [_∘_]
}
; isFunctor = record ; isFunctor = record
{ ident = funExt λ _ proj₂ ident { ident = funExt λ _ proj₂ ident
; distrib = funExt λ x sym assoc ; distrib = funExt λ x sym assoc
@ -73,8 +75,10 @@ Presheaf {' = '} = Functor (Opposite ) (Sets {'})
-- Alternate name: `yoneda` -- Alternate name: `yoneda`
presheaf : { ' : Level} { : Category '} Category.Object (Opposite ) Presheaf presheaf : { ' : Level} { : Category '} Category.Object (Opposite ) Presheaf
presheaf { = } B = record presheaf { = } B = record
{ func* = λ A .Arrow A B { raw = record
; func→ = λ f g ._∘_ g f { func* = λ A [ A , B ]
; func→ = λ f g [ g f ]
}
; isFunctor = record ; isFunctor = record
{ ident = funExt λ x proj₁ ident { ident = funExt λ x proj₁ ident
; distrib = funExt λ x assoc ; distrib = funExt λ x assoc

View file

@ -11,7 +11,7 @@ open import Data.Product renaming
) )
open import Data.Empty open import Data.Empty
import Function import Function
open import Cubical open import Cubical hiding (isSet)
open import Cubical.GradLemma using ( propIsEquiv ) open import Cubical.GradLemma using ( propIsEquiv )
∃! : {a b} {A : Set a} ∃! : {a b} {A : Set a}
@ -23,6 +23,9 @@ open import Cubical.GradLemma using ( propIsEquiv )
syntax ∃!-syntax (λ x B) = ∃![ x ] B syntax ∃!-syntax (λ x B) = ∃![ x ] B
IsSet : { : Level} (A : Set ) Set
IsSet A = {x y : A} (p q : x y) p q
record RawCategory ( ' : Level) : Set (lsuc (' )) where record RawCategory ( ' : Level) : Set (lsuc (' )) where
-- adding no-eta-equality can speed up type-checking. -- adding no-eta-equality can speed up type-checking.
-- ONLY IF you define your categories with copatterns though. -- ONLY IF you define your categories with copatterns though.
@ -50,16 +53,12 @@ record RawCategory ( ' : Level) : Set (lsuc (' ⊔ )) where
-- (univalent). -- (univalent).
record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc (a b)) where record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc (a b)) where
open RawCategory open RawCategory
-- (Object : Set )
-- (Arrow : Object → Object → Set ')
-- (𝟙 : {o : Object} → Arrow o o)
-- (_∘_ : { a b c : Object } → Arrow b c → Arrow a b → Arrow a c)
field field
assoc : {A B C D : Object} { f : Arrow A B } { g : Arrow B C } { h : Arrow C D } assoc : {A B C D : Object} { f : Arrow A B } { g : Arrow B C } { h : Arrow C D }
h (g f) (h g) f h (g f) (h g) f
ident : {A B : Object} {f : Arrow A B} ident : {A B : Object} {f : Arrow A B}
f 𝟙 f × 𝟙 f f f 𝟙 f × 𝟙 f f
arrow-is-set : {A B : Object} isSet (Arrow A B) arrowIsSet : {A B : Object} IsSet (Arrow A B)
Isomorphism : {A B} (f : Arrow A B) Set b Isomorphism : {A B} (f : Arrow A B) Set b
Isomorphism {A} {B} f = Σ[ g Arrow B A ] g f 𝟙 × f g 𝟙 Isomorphism {A} {B} f = Σ[ g Arrow B A ] g f 𝟙 × f g 𝟙
@ -73,11 +72,12 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
id-to-iso : (A B : Object) A B A B id-to-iso : (A B : Object) A B A B
id-to-iso A B eq = transp (\ i A eq i) (idIso A) id-to-iso A B eq = transp (\ i A eq i) (idIso A)
-- TODO: might want to implement isEquiv differently, there are 3 -- TODO: might want to implement isEquiv differently, there are 3
-- equivalent formulations in the book. -- equivalent formulations in the book.
Univalent : Set (a b)
Univalent = {A B : Object} isEquiv (A B) (A B) (id-to-iso A B)
field field
univalent : {A B : Object} isEquiv (A B) (A B) (id-to-iso A B) univalent : Univalent
module _ {A B : Object} where module _ {A B : Object} where
Epimorphism : {X : Object } (f : Arrow A B) Set b Epimorphism : {X : Object } (f : Arrow A B) Set b
@ -91,16 +91,15 @@ module _ {a} {b} { : RawCategory a b} where
-- This lemma will be useful to prove the equality of two categories. -- This lemma will be useful to prove the equality of two categories.
IsCategory-is-prop : isProp (IsCategory ) IsCategory-is-prop : isProp (IsCategory )
IsCategory-is-prop x y i = record IsCategory-is-prop x y i = record
{ assoc = x.arrow-is-set _ _ x.assoc y.assoc i { assoc = x.arrowIsSet x.assoc y.assoc i
; ident = ; ident =
( x.arrow-is-set _ _ (fst x.ident) (fst y.ident) i ( x.arrowIsSet (fst x.ident) (fst y.ident) i
, x.arrow-is-set _ _ (snd x.ident) (snd y.ident) i , x.arrowIsSet (snd x.ident) (snd y.ident) i
) )
-- ; arrow-is-set = {!λ x₁ y₁ p q → x.arrow-is-set _ _ p q!} ; arrowIsSet = λ p q
; arrow-is-set = λ _ _ p q
let let
golden : x.arrow-is-set _ _ p q y.arrow-is-set _ _ p q golden : x.arrowIsSet p q y.arrowIsSet p q
golden = λ j k l {!!} golden = {!!}
in in
golden i golden i
; univalent = λ y₁ {!!} ; univalent = λ y₁ {!!}
@ -109,148 +108,91 @@ module _ {a} {b} { : RawCategory a b} where
module x = IsCategory x module x = IsCategory x
module y = IsCategory y module y = IsCategory y
Category : (a b : Level) Set (lsuc (a b)) record Category (a b : Level) : Set (lsuc (a b)) where
Category a b = Σ (RawCategory a b) IsCategory
module Category {a b : Level} ( : Category a b) where
raw = fst
open RawCategory raw public
isCategory = snd
open RawCategory
-- _∈_ : ∀ {a b} ( : Category a b) → ( .fst .Object → Set b) → Set (a ⊔ b)
-- A ∈ =
Obj : {a b} Category a b Set a
Obj = .fst .Object
_[_,_] : { '} ( : Category ') (A : Obj ) (B : Obj ) Set '
[ A , B ] = .fst .Arrow A B
_[_∘_] : { '} ( : Category ') {A B C : Obj } (g : [ B , C ]) (f : [ A , B ]) [ A , C ]
[ g f ] = .fst ._∘_ g f
module _ { ' : Level} ( : Category ') {A B obj : Obj } where
IsProduct : (π₁ : [ obj , A ]) (π₂ : [ obj , B ]) Set ( ')
IsProduct π₁ π₂
= {X : Obj } (x₁ : [ X , A ]) (x₂ : [ X , B ])
∃![ x ] ( [ π₁ x ] x₁ × [ π₂ 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
-- field
-- isProduct : ∀ {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B)
-- → ∃![ x ] ( ._⊕_ π₁ x ≡ x₁ × . _⊕_ π₂ x ≡ x₂)
record Product { ' : Level} { : Category '} (A B : Obj ) : Set ( ') where
no-eta-equality
field field
obj : Obj raw : RawCategory a b
proj₁ : [ obj , A ] {{isCategory}} : IsCategory raw
proj₂ : [ obj , B ]
{{isProduct}} : IsProduct proj₁ proj₂
arrowProduct : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ]) private
[ X , obj ] module = RawCategory raw
arrowProduct π₁ π₂ = fst (isProduct π₁ π₂)
record HasProducts { ' : Level} ( : Category ') : Set ( ') where Object : Set a
field Object = .Object
product : (A B : Obj ) Product { = } A B
open Product Arrow = .Arrow
𝟙 = .𝟙
_∘_ = ._∘_
_[_,_] : (A : Object) (B : Object) Set b
_[_,_] = .Arrow
_[_∘_] : {A B C : Object} (g : .Arrow B C) (f : .Arrow A B) .Arrow A C
_[_∘_] = ._∘_
objectProduct : (A B : Obj ) Obj
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' : Obj } [ 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₂ ])
module _ {a b : Level} ( : Category a b) where module _ {a b : Level} ( : Category a b) where
private private
open Category open Category
module = RawCategory ( .fst)
OpRaw : RawCategory a b OpRaw : RawCategory a b
OpRaw = record RawCategory.Object OpRaw = Object
{ Object = .Object RawCategory.Arrow OpRaw = Function.flip Arrow
; Arrow = Function.flip .Arrow RawCategory.𝟙 OpRaw = 𝟙
; 𝟙 = .𝟙 RawCategory._∘_ OpRaw = Function.flip _∘_
; _∘_ = Function.flip (._∘_)
}
open IsCategory isCategory open IsCategory isCategory
OpIsCategory : IsCategory OpRaw OpIsCategory : IsCategory OpRaw
OpIsCategory = record IsCategory.assoc OpIsCategory = sym assoc
{ assoc = sym assoc IsCategory.ident OpIsCategory = swap ident
; ident = swap ident IsCategory.arrowIsSet OpIsCategory = arrowIsSet
; arrow-is-set = {!!} IsCategory.univalent OpIsCategory = {!!}
; univalent = {!!}
}
Opposite : Category a b Opposite : Category a b
Opposite = OpRaw , OpIsCategory raw Opposite = OpRaw
Category.isCategory Opposite = OpIsCategory
-- A consequence of no-eta-equality; `Opposite-is-involution` is no longer -- As demonstrated here a side-effect of having no-eta-equality on constructors
-- definitional - i.e.; you must match on the fields: -- means that we need to pick things apart to show that things are indeed
-- -- definitionally equal. I.e; a thing that would normally be provable in one
-- Opposite-is-involution : ∀ { '} → {C : Category {} {'}} → Opposite (Opposite C) ≡ C -- line now takes more than 20!!
-- Object (Opposite-is-involution {C = C} i) = Object C module _ {a b : Level} { : Category a b} where
-- Arrow (Opposite-is-involution i) = {!!}
-- 𝟙 (Opposite-is-involution i) = {!!}
-- _⊕_ (Opposite-is-involution i) = {!!}
-- assoc (Opposite-is-involution i) = {!!}
-- ident (Opposite-is-involution i) = {!!}
module _ { '} ( : Category ') {{hasProducts : HasProducts }} where
open HasProducts hasProducts
open Product hiding (obj)
private private
_×p_ : (A B : Obj ) Obj open RawCategory
_×p_ A B = Product.obj (product A B) module C = Category
rawOp : Category.raw (Opposite (Opposite )) Category.raw
Object (rawOp _) = C.Object
Arrow (rawOp _) = C.Arrow
𝟙 (rawOp _) = C.𝟙
_∘_ (rawOp _) = C._∘_
open Category
open IsCategory
module IsCat = IsCategory ( .isCategory)
rawIsCat : (i : I) IsCategory (rawOp i)
assoc (rawIsCat i) = IsCat.assoc
ident (rawIsCat i) = IsCat.ident
arrowIsSet (rawIsCat i) = IsCat.arrowIsSet
univalent (rawIsCat i) = IsCat.univalent
module _ (B C : Obj ) where Opposite-is-involution : Opposite (Opposite )
IsExponential : (Cᴮ : Obj ) [ Cᴮ ×p B , C ] Set ( ') raw (Opposite-is-involution i) = rawOp i
IsExponential Cᴮ eval = (A : Obj ) (f : [ A ×p B , C ]) isCategory (Opposite-is-involution i) = rawIsCat i
∃![ f~ ] ( [ eval parallelProduct f~ (Category.raw .𝟙)] f)
record Exponential : Set ( ') where
field
-- obj ≡ Cᴮ
obj : Obj
eval : [ 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?
exponentialsHaveProducts : HasProducts
exponentialsHaveProducts = hasProducts
transpose : (A : Obj ) [ A ×p B , C ] [ A , obj ]
transpose A f = fst (isExponential A f)
record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where
field
exponent : (A B : Obj ) Exponential A B
record CartesianClosed { ' : Level} ( : Category ') : Set ( ') where
field
{{hasProducts}} : HasProducts
{{hasExponentials}} : HasExponentials
module _ {a b : Level} ( : Category a b) where module _ {a b : Level} ( : Category a b) where
open Category
unique = isContr unique = isContr
IsInitial : Obj Set (a b) IsInitial : Object Set (a b)
IsInitial I = {X : Obj } unique ( [ I , X ]) IsInitial I = {X : Object } unique ( [ I , X ])
IsTerminal : Obj Set (a b) IsTerminal : Object Set (a b)
-- ∃![ ? ] ? -- ∃![ ? ] ?
IsTerminal T = {X : Obj } unique ( [ X , T ]) IsTerminal T = {X : Object } unique ( [ X , T ])
Initial : Set (a b) Initial : Set (a b)
Initial = Σ (Obj ) IsInitial Initial = Σ (Object ) IsInitial
Terminal : Set (a b) Terminal : Set (a b)
Terminal = Σ (Obj ) IsTerminal Terminal = Σ (Object ) IsTerminal

View file

@ -0,0 +1,12 @@
module Cat.Category.CartesianClosed where
open import Agda.Primitive
open import Cat.Category
open import Cat.Category.Product
open import Cat.Category.Exponential
record CartesianClosed { ' : Level} ( : Category ') : Set ( ') where
field
{{hasProducts}} : HasProducts
{{hasExponentials}} : HasExponentials

View file

@ -0,0 +1,39 @@
module Cat.Category.Exponential where
open import Agda.Primitive
open import Data.Product
open import Cubical
open import Cat.Category
open import Cat.Category.Product
open Category
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 : Object ) where
IsExponential : (Cᴮ : Object ) [ Cᴮ ×p B , C ] Set ( ')
IsExponential Cᴮ eval = (A : Object ) (f : [ A ×p B , C ])
∃![ f~ ] ( [ eval f~ |×| Category.𝟙 ] f)
record Exponential : Set ( ') where
field
-- obj ≡ Cᴮ
obj : Object
eval : [ 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?
exponentialsHaveProducts : HasProducts
exponentialsHaveProducts = hasProducts
transpose : (A : Object ) [ A ×p B , C ] [ A , obj ]
transpose A f = proj₁ (isExponential A f)
record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where
field
exponent : (A B : Object ) Exponential A B

View file

@ -1,42 +0,0 @@
{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Category.Free where
open import Agda.Primitive
open import Cubical hiding (Path)
open import Data.Product
open import Cat.Category as C
module _ { ' : Level} ( : Category ') where
private
open module = Category
postulate
Path : ( a b : Obj ) Set '
emptyPath : (o : Obj ) Path o o
concatenate : {a b c : Obj } Path b c Path a b Path a c
private
module _ {A B C D : Obj } {r : Path A B} {q : Path B C} {p : Path C D} where
postulate
p-assoc : concatenate {A} {C} {D} p (concatenate {A} {B} {C} q r)
concatenate {A} {B} {D} (concatenate {B} {C} {D} p q) r
module _ {A B : Obj } {p : Path A B} where
postulate
ident-r : concatenate {A} {A} {B} p (emptyPath A) p
ident-l : concatenate {A} {B} {B} (emptyPath B) p p
RawFree : RawCategory '
RawFree = record
{ Object = Obj
; Arrow = Path
; 𝟙 = λ {o} emptyPath o
; _∘_ = λ {a b c} concatenate {a} {b} {c}
}
RawIsCategoryFree : IsCategory RawFree
RawIsCategoryFree = record
{ assoc = p-assoc
; ident = ident-r , ident-l
; arrow-is-set = {!!}
; univalent = {!!}
}

View file

@ -0,0 +1,150 @@
module Cat.Category.Functor where
open import Agda.Primitive
open import Cubical
open import Function
open import Cat.Category
open Category hiding (_∘_ ; raw)
module _ {c c' d d'}
( : Category c c')
(𝔻 : Category d d')
where
private
= c c' d d'
𝓤 = Set
record RawFunctor : 𝓤 where
field
func* : Object Object 𝔻
func→ : {A B} [ A , B ] 𝔻 [ func* A , func* B ]
record IsFunctor (F : RawFunctor) : 𝓤 where
open RawFunctor F
field
ident : {c : Object } func→ (𝟙 {c}) 𝟙 𝔻 {func* c}
distrib : {A B C : Object } {f : [ A , B ]} {g : [ B , C ]}
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ]
record Functor : Set (c c' d d') where
field
raw : RawFunctor
{{isFunctor}} : IsFunctor raw
private
module R = RawFunctor raw
func* : Object Object 𝔻
func* = R.func*
func→ : {A B} [ A , B ] 𝔻 [ func* A , func* B ]
func→ = R.func→
open IsFunctor
open Functor
-- TODO: Is `IsFunctor` a proposition?
module _
{a b : Level}
{ 𝔻 : Category a b}
{F : RawFunctor 𝔻}
where
private
module 𝔻 = IsCategory (isCategory 𝔻)
-- isProp : Set
-- isProp = (x y : A) → x ≡ y
IsFunctorIsProp : isProp (IsFunctor _ _ F)
IsFunctorIsProp isF0 isF1 i = record
{ ident = 𝔻.arrowIsSet isF0.ident isF1.ident i
; distrib = 𝔻.arrowIsSet isF0.distrib isF1.distrib i
}
where
module isF0 = IsFunctor isF0
module isF1 = IsFunctor isF1
-- Alternate version of above where `F` is a path in
module _
{a b : Level}
{ 𝔻 : Category a b}
{F : I RawFunctor 𝔻}
where
private
module 𝔻 = IsCategory (isCategory 𝔻)
IsProp' : { : Level} (A : I Set ) Set
IsProp' A = (a0 : A i0) (a1 : A i1) A [ a0 a1 ]
postulate IsFunctorIsProp' : IsProp' λ i IsFunctor _ _ (F i)
-- IsFunctorIsProp' isF0 isF1 i = record
-- { ident = {!𝔻.arrowIsSet {!isF0.ident!} {!isF1.ident!} i!}
-- ; distrib = {!𝔻.arrowIsSet {!isF0.distrib!} {!isF1.distrib!} i!}
-- }
-- where
-- module isF0 = IsFunctor isF0
-- module isF1 = IsFunctor isF1
module _ { ' : Level} { 𝔻 : Category '} where
Functor≡ : {F G : Functor 𝔻}
(eq* : func* F func* G)
(eq→ : (λ i {x y} [ x , y ] 𝔻 [ eq* i x , eq* i y ])
[ func→ F func→ G ])
F G
Functor≡ {F} {G} eq* eq→ i = record
{ raw = eqR i
; isFunctor = f i
}
where
eqR : raw F raw G
eqR i = record { func* = eq* i ; func→ = eq→ i }
postulate T : isSet (IsFunctor _ _ (raw F))
f : (λ i IsFunctor 𝔻 (eqR i)) [ isFunctor F isFunctor G ]
f = IsFunctorIsProp' (isFunctor F) (isFunctor G)
module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where
private
F* = func* F
F→ = func→ F
G* = func* G
G→ = func→ G
module _ {a0 a1 a2 : Object A} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where
dist : (F→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ]
dist = begin
(F→ G→) (A [ α1 α0 ]) ≡⟨ refl
F→ (G→ (A [ α1 α0 ])) ≡⟨ cong F→ (G .isFunctor .distrib)
F→ (B [ G→ α1 G→ α0 ]) ≡⟨ F .isFunctor .distrib
C [ (F→ G→) α1 (F→ G→) α0 ]
_∘fr_ : RawFunctor A C
RawFunctor.func* _∘fr_ = F* G*
RawFunctor.func→ _∘fr_ = F→ G→
instance
isFunctor' : IsFunctor A C _∘fr_
isFunctor' = record
{ ident = begin
(F→ G→) (𝟙 A) ≡⟨ refl
F→ (G→ (𝟙 A)) ≡⟨ cong F→ (G .isFunctor .ident)
F→ (𝟙 B) ≡⟨ F .isFunctor .ident
𝟙 C
; distrib = dist
}
_∘f_ : Functor A C
raw _∘f_ = _∘fr_
-- The identity functor
identity : { '} {C : Category '} Functor C C
identity = record
{ raw = record
{ func* = λ x x
; func→ = λ x x
}
; isFunctor = record
{ ident = refl
; distrib = refl
}
}

View file

@ -0,0 +1,54 @@
module Cat.Category.Product where
open import Agda.Primitive
open import Cubical
open import Data.Product as P hiding (_×_)
open import Cat.Category
open Category
module _ { ' : Level} ( : Category ') {A B obj : Object } where
IsProduct : (π₁ : [ obj , A ]) (π₂ : [ obj , B ]) Set ( ')
IsProduct π₁ π₂
= {X : Object } (x₁ : [ X , A ]) (x₂ : [ X , B ])
∃![ x ] ( [ π₁ x ] x₁ P.× [ π₂ x ] x₂)
-- Tip from Andrea; Consider this style for efficiency:
-- record IsProduct {a b : Level} ( : Category a b)
-- {A B obj : Object } (π₁ : Arrow obj A) (π₂ : Arrow obj B) : Set (a ⊔ b) where
-- field
-- 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
field
obj : Object
proj₁ : [ obj , A ]
proj₂ : [ obj , B ]
{{isProduct}} : IsProduct proj₁ proj₂
_P[_×_] : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
[ X , obj ]
_P[_×_] π₁ π₂ = proj₁ (isProduct π₁ π₂)
record HasProducts { ' : Level} ( : Category ') : Set ( ') where
field
product : (A B : Object ) Product { = } A B
open Product
_×_ : (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
_|×|_ : {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

@ -7,12 +7,12 @@ open import Data.Product
open import Cubical open import Cubical
open import Cat.Category open import Cat.Category
open import Cat.Functor open import Cat.Category.Functor
open import Cat.Categories.Sets open import Cat.Categories.Sets
open import Cat.Equality open import Cat.Equality
open Equality.Data.Product open Equality.Data.Product
module _ { ' : Level} { : Category '} { A B : .Category.Object } {X : .Category.Object} (f : .Category.Arrow A B) where module _ { ' : Level} { : Category '} { A B : Category.Object } {X : Category.Object } (f : Category.Arrow A B) where
open Category open Category
open IsCategory (isCategory) open IsCategory (isCategory)
@ -51,7 +51,6 @@ epi-mono-is-not-iso f =
open import Cat.Category open import Cat.Category
open Category open Category
open import Cat.Functor
open Functor open Functor
-- module _ { : Level} { : Category } -- module _ { : Level} { : Category }
@ -61,7 +60,7 @@ open Functor
-- open import Cat.Categories.Fun -- open import Cat.Categories.Fun
-- open import Cat.Categories.Sets -- open import Cat.Categories.Sets
-- -- module Cat = Cat.Categories.Cat -- -- module Cat = Cat.Categories.Cat
-- open Exponential -- open import Cat.Category.Exponential
-- private -- private
-- Cat = Cat -- Cat = Cat
-- prshf = presheaf { = } -- prshf = presheaf { = }

View file

@ -4,7 +4,7 @@ open import Agda.Primitive
open import Data.Product open import Data.Product
open import Cat.Category open import Cat.Category
open import Cat.Functor open import Cat.Category.Functor
open import Cat.Categories.Fam open import Cat.Categories.Fam
open Category open Category
@ -17,25 +17,27 @@ module _ {a b : Level} where
-- "A base category" -- "A base category"
: Category a b : Category a b
-- It's objects are called contexts -- It's objects are called contexts
Contexts = .Object Contexts = Object
-- It's arrows are called substitutions -- It's arrows are called substitutions
Substitutions = .Arrow Substitutions = Arrow
field field
-- A functor T -- A functor T
T : Functor (Opposite ) (Fam a b) T : Functor (Opposite ) (Fam a b)
-- Empty context -- Empty context
[] : Terminal [] : Terminal
Type : (Γ : .Object) Set a private
Type Γ = proj₁ (T .func* Γ) module T = Functor T
Type : (Γ : Object ) Set a
Type Γ = proj₁ (T.func* Γ)
module _ {Γ : .Object} {A : Type Γ} where module _ {Γ : Object } {A : Type Γ} where
module _ {A B : .Object} {γ : [ A , B ]} where module _ {A B : Object } {γ : [ A , B ]} where
k : Σ (proj₁ (func* T B) proj₁ (func* T A)) k : Σ (proj₁ (func* T B) proj₁ (func* T A))
(λ f (λ f
{x : proj₁ (func* T B)} {x : proj₁ (func* T B)}
proj₂ (func* T B) x proj₂ (func* T A) (f x)) proj₂ (func* T B) x proj₂ (func* T A) (f x))
k = T .func→ γ k = T.func→ γ
k₁ : proj₁ (func* T B) proj₁ (func* T A) k₁ : proj₁ (func* T B) proj₁ (func* T A)
k₁ = proj₁ k k₁ = proj₁ k
k₂ : ({x : proj₁ (func* T B)} k₂ : ({x : proj₁ (func* T B)}
@ -44,8 +46,8 @@ module _ {a b : Level} where
record ContextComprehension : Set (a b) where record ContextComprehension : Set (a b) where
field field
Γ&A : .Object Γ&A : Object
proj1 : .Arrow Γ&A Γ proj1 : [ Γ&A , Γ ]
-- proj2 : ???? -- proj2 : ????
-- if γ : [ A , B ] -- if γ : [ A , B ]

View file

@ -1,97 +0,0 @@
module Cat.Functor where
open import Agda.Primitive
open import Cubical
open import Function
open import Cat.Category
open Category hiding (_∘_)
module _ {c c' d d'} ( : Category c c') (𝔻 : Category d d') where
record IsFunctor
(func* : Obj Obj 𝔻)
(func→ : {A B : Obj } [ A , B ] 𝔻 [ func* A , func* B ])
: Set (c c' d d') where
field
ident : {c : Obj } func→ ( .𝟙 {c}) 𝔻 .𝟙 {func* c}
-- TODO: Avoid use of ugly explicit arguments somehow.
-- This guy managed to do it:
-- https://github.com/copumpkin/categories/blob/master/Categories/Functor/Core.agda
distrib : {A B C : .Object} {f : [ A , B ]} {g : [ B , C ]}
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ]
record Functor : Set (c c' d d') where
field
func* : .Object 𝔻 .Object
func→ : {A B} [ A , B ] 𝔻 [ func* A , func* B ]
{{isFunctor}} : IsFunctor func* func→
open IsFunctor
open Functor
module _ { ' : Level} { 𝔻 : Category '} where
IsFunctor≡
: {func* : .Object 𝔻 .Object}
{func→ : {A B : .Object} .Arrow A B 𝔻 .Arrow (func* A) (func* B)}
{F G : IsFunctor 𝔻 func* func→}
(eqI
: (λ i {A} func→ ( .𝟙 {A}) 𝔻 .𝟙 {func* A})
[ F .ident G .ident ])
(eqD :
(λ i {A B C} {f : [ A , B ]} {g : [ B , C ]}
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ])
[ F .distrib G .distrib ])
(λ _ IsFunctor 𝔻 (λ i func* i) func→) [ F G ]
IsFunctor≡ eqI eqD i = record { ident = eqI i ; distrib = eqD i }
Functor≡ : {F G : Functor 𝔻}
(eq* : F .func* G .func*)
(eq→ : (λ i {x y} [ x , y ] 𝔻 [ eq* i x , eq* i y ])
[ F .func→ G .func→ ])
-- → (eqIsF : PathP (λ i → IsFunctor 𝔻 (eq* i) (eq→ i)) (F .isFunctor) (G .isFunctor))
(eqIsFunctor : (λ i IsFunctor 𝔻 (eq* i) (eq→ i)) [ F .isFunctor G .isFunctor ])
F G
Functor≡ eq* eq→ eqIsFunctor i = record { func* = eq* i ; func→ = eq→ i ; isFunctor = eqIsFunctor i }
module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where
private
F* = F .func*
F→ = F .func→
G* = G .func*
G→ = G .func→
module _ {a0 a1 a2 : A .Object} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where
dist : (F→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ]
dist = begin
(F→ G→) (A [ α1 α0 ]) ≡⟨ refl
F→ (G→ (A [ α1 α0 ])) ≡⟨ cong F→ (G .isFunctor .distrib)
F→ (B [ G→ α1 G→ α0 ]) ≡⟨ F .isFunctor .distrib
C [ (F→ G→) α1 (F→ G→) α0 ]
_∘f_ : Functor A C
_∘f_ =
record
{ func* = F* G*
; func→ = F→ G→
; isFunctor = record
{ ident = begin
(F→ G→) (A .𝟙) ≡⟨ refl
F→ (G→ (A .𝟙)) ≡⟨ cong F→ (G .isFunctor .ident)
F→ (B .𝟙) ≡⟨ F .isFunctor .ident
C .𝟙
; distrib = dist
}
}
-- The identity functor
identity : { '} {C : Category '} Functor C C
identity = record
{ func* = λ x x
; func→ = λ x x
; isFunctor = record
{ ident = refl
; distrib = refl
}
}