Merge branch 'dev'

This commit is contained in:
Frederik Hanghøj Iversen 2018-02-23 12:57:19 +01:00
commit 4fac927613
12 changed files with 368 additions and 459 deletions

View file

@ -1,6 +1,23 @@
Changelog Changelog
========= =========
Version 1.2.0
-------------
This version is mainly a huge refactor.
I've renamed
* `distrib` to `isDistributive`
* `arrowIsSet` to `arrowsAreSets`
* `ident` to `isIdentity`
* `assoc` to `isAssociative`
And added "type-synonyms" for all of these. Their names should now match their
type. So e.g. `isDistributive` has type `IsDistributive`.
I've also changed how names are exported in `Functor` to be in line with
`Category`.
Version 1.1.0 Version 1.1.0
------------- -------------
In this version categories have been refactored - there's now a notion of a raw In this version categories have been refactored - there's now a notion of a raw

View file

@ -16,107 +16,58 @@ open import Cat.Category.Exponential
open import Cat.Equality open import Cat.Equality
open Equality.Data.Product open Equality.Data.Product
open Functor open Functor using (func→ ; func*)
open IsFunctor open Category using (Object ; 𝟙)
open Category hiding (_∘_)
-- The category of categories -- The category of categories
module _ ( ' : Level) where module _ ( ' : Level) where
private private
module _ {𝔸 𝔹 𝔻 : Category '} {F : Functor 𝔸 𝔹} {G : Functor 𝔹 } {H : Functor 𝔻} where module _ {𝔸 𝔹 𝔻 : Category '} {F : Functor 𝔸 𝔹} {G : Functor 𝔹 } {H : Functor 𝔻} where
private
eq* : func* (H ∘f (G ∘f F)) func* ((H ∘f G) ∘f F)
eq* = refl
eq→ : PathP
(λ i {A B : Object 𝔸} 𝔸 [ A , B ] 𝔻 [ eq* i A , eq* i B ])
(func→ (H ∘f (G ∘f F))) (func→ ((H ∘f G) ∘f F))
eq→ = refl
postulate
eqI
: (λ i {A : Object 𝔸} eq→ i (𝟙 𝔸 {A}) 𝟙 𝔻 {eq* i A})
[ (H ∘f (G ∘f F)) .isFunctor .ident
((H ∘f G) ∘f F) .isFunctor .ident
]
eqD
: (λ i {A B C} {f : 𝔸 [ A , B ]} {g : 𝔸 [ B , C ]}
eq→ i (𝔸 [ g f ]) 𝔻 [ eq→ i g eq→ i f ])
[ (H ∘f (G ∘f F)) .isFunctor .distrib
((H ∘f G) ∘f F) .isFunctor .distrib
]
assc : H ∘f (G ∘f F) (H ∘f G) ∘f F assc : H ∘f (G ∘f F) (H ∘f G) ∘f F
assc = Functor≡ eq* eq→ (IsFunctor≡ eqI eqD) assc = Functor≡ refl refl
module _ { 𝔻 : Category '} {F : Functor 𝔻} where module _ { 𝔻 : Category '} {F : Functor 𝔻} where
module _ where ident-r : F ∘f identity F
private ident-r = Functor≡ refl refl
eq* : (func* F) (func* (identity {C = })) func* F
eq* = refl
-- lemmm : func→ {C = A} {D = B} (f ∘f identity) ≡ func→ f
eq→ : PathP
(λ i
{x y : Object } Arrow x y Arrow 𝔻 (func* F x) (func* F y))
(func→ (F ∘f identity)) (func→ F)
eq→ = refl
postulate
eqI-r
: (λ i {c : Object } (λ _ 𝔻 [ func* F c , func* F c ])
[ func→ F (𝟙 ) 𝟙 𝔻 ])
[(F ∘f identity) .isFunctor .ident F .isFunctor .ident ]
eqD-r : PathP
(λ i
{A B C : Object } {f : [ 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
ident-r = Functor≡ eq* eq→ (IsFunctor≡ eqI-r eqD-r)
module _ where
private
postulate
eq* : (identity ∘f F) .func* F .func*
eq→ : PathP
(λ i {x y : Object } [ x , y ] 𝔻 [ eq* i x , eq* i y ])
((identity ∘f F) .func→) (F .func→)
eqI : (λ i {A : Object } eq→ i (𝟙 {A}) 𝟙 𝔻 {eq* i A})
[ ((identity ∘f F) .isFunctor .ident) (F .isFunctor .ident) ]
eqD : PathP (λ i {A B C : Object } {f : [ 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 }
RawCat : RawCategory (lsuc ( ')) ( ') ident-l : identity ∘f F F
RawCat = ident-l = Functor≡ refl refl
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 ( ')) ( ') RawCat : RawCategory (lsuc ( ')) ( ')
raw Cat = RawCat RawCat =
record
{ Object = Category '
; Arrow = Functor
; 𝟙 = identity
; _∘_ = _∘f_
}
private
open RawCategory RawCat
isAssociative : IsAssociative
isAssociative {f = F} {G} {H} = assc {F = F} {G = G} {H = H}
-- TODO: Rename `ident'` to `ident` after changing how names are exposed in Functor.
ident' : IsIdentity identity
ident' = ident-r , ident-l
-- NB! `ArrowsAreSets RawCat` is *not* provable. The type of functors,
-- however, form a groupoid! Therefore there is no (1-)category of
-- categories. There does, however, exist a 2-category of 1-categories.
module _ { ' : Level} where -- Because of the note above there is not category of categories.
Cat : (unprovable : IsCategory RawCat) Category (lsuc ( ')) ( ')
Category.raw (Cat _) = RawCat
Category.isCategory (Cat unprovable) = unprovable
-- Category.raw Cat _ = RawCat
-- Category.isCategory Cat unprovable = unprovable
-- The following to some extend depends on the category of categories being a
-- category. In some places it may not actually be needed, however.
module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
module _ ( 𝔻 : Category ') where module _ ( 𝔻 : Category ') where
private private
Catt = Cat ' Catt = Cat ' unprovable
: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') = [ c , c' ] × 𝔻 [ d , d' ]
:𝟙: : {o : :Object:} :Arrow: o o :𝟙: : {o : :Object:} :Arrow: o o
:𝟙: = 𝟙 , 𝟙 𝔻 :𝟙: = 𝟙 , 𝟙 𝔻
_:⊕:_ : _:⊕:_ :
@ -131,70 +82,67 @@ module _ { ' : Level} where
RawCategory.Arrow :rawProduct: = :Arrow: RawCategory.Arrow :rawProduct: = :Arrow:
RawCategory.𝟙 :rawProduct: = :𝟙: RawCategory.𝟙 :rawProduct: = :𝟙:
RawCategory._∘_ :rawProduct: = _:⊕:_ RawCategory._∘_ :rawProduct: = _:⊕:_
open RawCategory :rawProduct:
module C = IsCategory ( .isCategory) module C = Category
module D = IsCategory (𝔻 .isCategory) module D = Category 𝔻
postulate open import Cubical.Sigma
issSet : {A B : RawCategory.Object :rawProduct:} isSet (RawCategory.Arrow :rawProduct: A B) issSet : {A B : RawCategory.Object :rawProduct:} isSet (Arrow A B)
issSet = setSig {sA = C.arrowsAreSets} {sB = λ x D.arrowsAreSets}
ident' : IsIdentity :𝟙:
ident'
= Σ≡ (fst C.isIdentity) (fst D.isIdentity)
, Σ≡ (snd C.isIdentity) (snd D.isIdentity)
postulate univalent : Univalence.Univalent :rawProduct: ident'
instance instance
:isCategory: : IsCategory :rawProduct: :isCategory: : IsCategory :rawProduct:
-- :isCategory: = record IsCategory.isAssociative :isCategory: = Σ≡ C.isAssociative D.isAssociative
-- { assoc = Σ≡ C.assoc D.assoc IsCategory.isIdentity :isCategory: = ident'
-- ; ident IsCategory.arrowsAreSets :isCategory: = issSet
-- = Σ≡ (fst C.ident) (fst D.ident) IsCategory.univalent :isCategory: = univalent
-- , Σ≡ (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)
IsCategory.arrow-is-set :isCategory: = issSet
IsCategory.univalent :isCategory: = {!!}
:product: : Category ' :product: : Category '
raw :product: = :rawProduct: Category.raw :product: = :rawProduct:
proj₁ : Catt [ :product: , ] proj₁ : Catt [ :product: , ]
proj₁ = record { func* = fst ; func→ = fst ; isFunctor = record { ident = refl ; distrib = refl } } proj₁ = record
{ raw = record { func* = fst ; func→ = fst }
; isFunctor = record { isIdentity = refl ; isDistributive = refl }
}
proj₂ : Catt [ :product: , 𝔻 ] proj₂ : Catt [ :product: , 𝔻 ]
proj₂ = record { func* = snd ; func→ = snd ; isFunctor = record { ident = refl ; distrib = refl } } proj₂ = record
{ raw = record { func* = snd ; func→ = snd }
; isFunctor = record { isIdentity = refl ; isDistributive = refl }
}
module _ {X : Object Catt} (x₁ : Catt [ X , ]) (x₂ : Catt [ X , 𝔻 ]) where module _ {X : Object Catt} (x₁ : Catt [ X , ]) (x₂ : Catt [ X , 𝔻 ]) where
open Functor x : Functor X :product:
x = record
{ raw = record
{ func* = λ x x₁ .func* x , x₂ .func* x
; func→ = λ x func→ x₁ x , func→ x₂ x
}
; isFunctor = record
{ isIdentity = Σ≡ x₁.isIdentity x₂.isIdentity
; isDistributive = Σ≡ x₁.isDistributive x₂.isDistributive
}
}
where
open module x = Functor x₁
open module x = Functor x₂
postulate x : Functor X :product: isUniqL : Catt [ proj₁ x ] x₁
-- x = record isUniqL = Functor≡ eq* eq→
-- { func* = λ x → x₁ .func* x , x₂ .func* x where
-- ; func→ = λ x → func→ x₁ x , func→ x₂ x eq* : (Catt [ proj₁ x ]) .func* x₁ .func*
-- ; isFunctor = record eq* = refl
-- { ident = Σ≡ x₁.ident x₂.ident eq→ : (λ i {A : Object X} {B : Object X} X [ A , B ] [ eq* i A , eq* i B ])
-- ; distrib = Σ≡ x₁.distrib x₂.distrib [ (Catt [ proj₁ x ]) .func→ x₁ .func→ ]
-- } eq→ = refl
-- }
-- where
-- open module x₁ = IsFunctor (x₁ .isFunctor)
-- open module x₂ = IsFunctor (x₂ .isFunctor)
-- Turned into postulate after: isUniqR : Catt [ proj₂ x ] x₂
-- > commit e8215b2c051062c6301abc9b3f6ec67106259758 (HEAD -> dev, github/dev) isUniqR = Functor≡ refl refl
-- > 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₂
-- isUniqR = Functor≡ refl refl {!!} {!!}
isUniq : Catt [ proj₁ x ] x₁ × Catt [ proj₂ x ] x₂ isUniq : Catt [ proj₁ x ] x₁ × Catt [ proj₂ x ] x₂
isUniq = isUniqL , isUniqR isUniq = isUniqL , isUniqR
@ -203,36 +151,37 @@ module _ { ' : Level} where
uniq = x , isUniq uniq = x , isUniq
instance instance
isProduct : IsProduct (Cat ') proj₁ proj₂ isProduct : IsProduct Catt proj₁ proj₂
isProduct = uniq isProduct = uniq
product : Product { = (Cat ')} 𝔻 product : Product { = Catt} 𝔻
product = record product = record
{ obj = :product: { obj = :product:
; proj₁ = proj₁ ; proj₁ = proj₁
; proj₂ = proj₂ ; proj₂ = proj₂
} }
module _ { ' : Level} where module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
Catt = Cat ' unprovable
instance instance
hasProducts : HasProducts (Cat ') hasProducts : HasProducts Catt
hasProducts = record { product = product } hasProducts = record { product = product unprovable }
-- Basically proves that `Cat ` is cartesian closed. -- Basically proves that `Cat ` is cartesian closed.
module _ ( : Level) where module _ ( : Level) (unprovable : IsCategory (RawCat )) where
private private
open Data.Product open Data.Product
open import Cat.Categories.Fun open import Cat.Categories.Fun
Cat : Category (lsuc ( )) ( ) Cat : Category (lsuc ( )) ( )
Cat = Cat Cat = Cat unprovable
module _ ( 𝔻 : Category ) where module _ ( 𝔻 : Category ) where
private private
:obj: : Object (Cat ) :obj: : Object Cat
:obj: = Fun { = } {𝔻 = 𝔻} :obj: = Fun { = } {𝔻 = 𝔻}
:func*: : Functor 𝔻 × Object Object 𝔻 :func*: : Functor 𝔻 × Object Object 𝔻
:func*: (F , A) = F .func* A :func*: (F , A) = func* F A
module _ {dom cod : Functor 𝔻 × Object } where module _ {dom cod : Functor 𝔻 × Object } where
private private
@ -247,30 +196,30 @@ module _ ( : Level) where
B = proj₂ cod B = proj₂ cod
:func→: : (pobj : NaturalTransformation F G × [ A , B ]) :func→: : (pobj : NaturalTransformation F G × [ A , B ])
𝔻 [ F .func* A , G .func* B ] 𝔻 [ func* F A , func* G B ]
:func→: ((θ , θNat) , f) = result :func→: ((θ , θNat) , f) = result
where where
θA : 𝔻 [ F .func* A , G .func* A ] θA : 𝔻 [ func* F A , func* G A ]
θA = θ A θA = θ A
θB : 𝔻 [ F .func* B , G .func* B ] θB : 𝔻 [ func* F B , func* G B ]
θB = θ B θB = θ B
F→f : 𝔻 [ F .func* A , F .func* B ] F→f : 𝔻 [ func* F A , func* F B ]
F→f = F .func→ f F→f = func→ F f
G→f : 𝔻 [ G .func* A , G .func* B ] G→f : 𝔻 [ func* G A , func* G B ]
G→f = G .func→ f G→f = func→ G f
l : 𝔻 [ F .func* A , G .func* B ] l : 𝔻 [ func* F A , func* G B ]
l = 𝔻 [ θB F→f ] l = 𝔻 [ θB F→f ]
r : 𝔻 [ F .func* A , G .func* B ] r : 𝔻 [ func* F A , func* G 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 : 𝔻 [ F .func* A , G .func* B ] result : 𝔻 [ func* F A , func* G B ]
result = l result = l
_×p_ = product _×p_ = product unprovable
module _ {c : Functor 𝔻 × Object } where module _ {c : Functor 𝔻 × Object } where
private private
@ -281,21 +230,21 @@ module _ ( : Level) where
-- NaturalTransformation F G × .Arrow A B -- NaturalTransformation F G × .Arrow A B
-- :ident: : :func→: {c} {c} (identityNat F , .𝟙) 𝔻 .𝟙 -- :ident: : :func→: {c} {c} (identityNat F , .𝟙) 𝔻 .𝟙
-- :ident: = trans (proj₂ 𝔻.ident) (F .ident) -- :ident: = trans (proj₂ 𝔻.isIdentity) (F .isIdentity)
-- 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 , 𝟙 {A = proj₂ c}) 𝟙 𝔻
:ident: = begin :ident: = begin
:func→: {c} {c} (𝟙 (Product.obj (:obj: ×p )) {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 func→ F (𝟙 )] ≡⟨⟩
𝔻 [ 𝟙 𝔻 F .func→ (𝟙 )] ≡⟨ proj₂ 𝔻.ident 𝔻 [ 𝟙 𝔻 func→ F (𝟙 )] ≡⟨ proj₂ 𝔻.isIdentity
F .func→ (𝟙 ) ≡⟨ F.ident func→ F (𝟙 ) ≡⟨ F.isIdentity
𝟙 𝔻 𝟙 𝔻
where where
open module 𝔻 = IsCategory (𝔻 .isCategory) open module 𝔻 = Category 𝔻
open module F = IsFunctor (F .isFunctor) open module F = Functor F
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₁
@ -330,48 +279,51 @@ module _ ( : Level) where
ηθ = proj₁ ηθNT ηθ = proj₁ ηθNT
ηθNat = proj₂ ηθNT ηθNat = proj₂ ηθNT
:distrib: : :isDistributive: :
𝔻 [ 𝔻 [ η C θ C ] F .func→ ( [ g f ] ) ] 𝔻 [ 𝔻 [ η C θ C ] func→ F ( [ g f ] ) ]
𝔻 [ 𝔻 [ η C G .func→ g ] 𝔻 [ θ B F .func→ f ] ] 𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ θ B func→ F f ] ]
:distrib: = begin :isDistributive: = begin
𝔻 [ (ηθ C) F .func→ ( [ g f ]) ] 𝔻 [ (ηθ C) func→ F ( [ g f ]) ]
≡⟨ ηθNat ( [ g f ]) ≡⟨ ηθNat ( [ g f ])
𝔻 [ H .func→ ( [ g f ]) (ηθ A) ] 𝔻 [ func→ H ( [ g f ]) (ηθ A) ]
≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.distrib) ≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.isDistributive)
𝔻 [ 𝔻 [ H .func→ g H .func→ f ] (ηθ A) ] 𝔻 [ 𝔻 [ func→ H g func→ H f ] (ηθ A) ]
≡⟨ sym assoc ≡⟨ sym isAssociative
𝔻 [ H .func→ g 𝔻 [ H .func→ f ηθ A ] ] 𝔻 [ func→ H g 𝔻 [ func→ H f ηθ A ] ]
≡⟨ cong (λ φ 𝔻 [ H .func→ g φ ]) assoc ≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) isAssociative
𝔻 [ H .func→ g 𝔻 [ 𝔻 [ H .func→ f η A ] θ A ] ] 𝔻 [ func→ H g 𝔻 [ 𝔻 [ func→ H f η A ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ H .func→ g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f))) ≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f)))
𝔻 [ H .func→ g 𝔻 [ 𝔻 [ η B G .func→ f ] θ A ] ] 𝔻 [ func→ H g 𝔻 [ 𝔻 [ η B func→ G f ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ H .func→ g φ ]) (sym assoc) ≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) (sym isAssociative)
𝔻 [ H .func→ g 𝔻 [ η B 𝔻 [ G .func→ f θ A ] ] ] ≡⟨ assoc 𝔻 [ func→ H g 𝔻 [ η B 𝔻 [ func→ G f θ A ] ] ]
𝔻 [ 𝔻 [ H .func→ g η B ] 𝔻 [ G .func→ f θ A ] ] ≡⟨ isAssociative
≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ G .func→ f θ A ] ]) (sym (ηNat g)) 𝔻 [ 𝔻 [ func→ H g η B ] 𝔻 [ func→ G f θ A ] ]
𝔻 [ 𝔻 [ η C G .func→ g ] 𝔻 [ G .func→ f θ A ] ] ≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ func→ G f θ A ] ]) (sym (ηNat g))
≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C G .func→ g ] φ ]) (sym (θNat f)) 𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ func→ G f θ A ] ]
𝔻 [ 𝔻 [ η C G .func→ g ] 𝔻 [ θ B F .func→ f ] ] ≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C func→ G g ] φ ]) (sym (θNat f))
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ θ B func→ F f ] ]
where where
open IsCategory (𝔻 .isCategory) open Category 𝔻
open module H = IsFunctor (H .isFunctor) module H = Functor H
:eval: : Functor ((:obj: ×p ) .Product.obj) 𝔻 :eval: : Functor ((:obj: ×p ) .Product.obj) 𝔻
:eval: = record :eval: = record
{ func* = :func*: { raw = record
; func→ = λ {dom} {cod} :func→: {dom} {cod} { func* = :func*:
; func→ = λ {dom} {cod} :func→: {dom} {cod}
}
; isFunctor = record ; isFunctor = record
{ ident = λ {o} :ident: {o} { isIdentity = λ {o} :ident: {o}
; distrib = λ {f u n k y} :distrib: {f} {u} {n} {k} {y} ; isDistributive = λ {f u n k y} :isDistributive: {f} {u} {n} {k} {y}
} }
} }
module _ (𝔸 : Category ) (F : Functor ((𝔸 ×p ) .Product.obj) 𝔻) where module _ (𝔸 : Category ) (F : Functor ((𝔸 ×p ) .Product.obj) 𝔻) where
open HasProducts (hasProducts {} {}) renaming (_|×|_ to parallelProduct) open HasProducts (hasProducts {} {} unprovable) 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 {A = })) ] F
-- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (𝟙 Cat {o = })) ] ≡ F -- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (𝟙 Cat {o = })) ] ≡ F
-- eq' : (Cat [ :eval: ∘ -- eq' : (Cat [ :eval: ∘
-- (record { product = product } HasProducts.|×| transpose) -- (record { product = product } HasProducts.|×| transpose)
@ -384,10 +336,11 @@ module _ ( : Level) where
-- :eval: (parallelProduct F~ (𝟙 Cat {o = }))] F) catTranspose = -- :eval: (parallelProduct F~ (𝟙 Cat {o = }))] F) catTranspose =
-- transpose , eq -- transpose , eq
:isExponential: : IsExponential Cat 𝔻 :obj: :eval: postulate :isExponential: : IsExponential Cat 𝔻 :obj: :eval:
:isExponential: = {!catTranspose!} -- :isExponential: : IsExponential Cat 𝔻 :obj: :eval:
where -- :isExponential: = {!catTranspose!}
open HasProducts (hasProducts {} {}) using (_|×|_) -- where
-- open HasProducts (hasProducts {} {} unprovable) using (_|×|_)
-- :isExponential: = λ 𝔸 F transpose 𝔸 F , eq' 𝔸 F -- :isExponential: = λ 𝔸 F transpose 𝔸 F , eq' 𝔸 F
-- :exponent: : Exponential (Cat ) A B -- :exponent: : Exponential (Cat ) A B
@ -398,5 +351,5 @@ module _ ( : Level) where
; isExponential = :isExponential: ; isExponential = :isExponential:
} }
hasExponentials : HasExponentials (Cat ) hasExponentials : HasExponentials Cat
hasExponentials = record { exponent = :exponent: } hasExponentials = record { exponent = :exponent: }

View file

@ -25,12 +25,12 @@ module _ (a b : Level) where
c g f = _∘_ {c = c} g f c g f = _∘_ {c = c} g f
module _ {A B C D : Obj'} {f : Arr A B} {g : Arr B C} {h : Arr C D} where module _ {A B C D : Obj'} {f : Arr A B} {g : Arr B C} {h : Arr C D} where
assoc : (D h C g f ) D D h g f isAssociative : (D h C g f ) D D h g f
assoc = Σ≡ refl refl isAssociative = Σ≡ refl refl
module _ {A B : Obj'} {f : Arr A B} where module _ {A B : Obj'} {f : Arr A B} where
ident : B f one f × B one {B} f f isIdentity : B f one f × B one {B} f f
ident = (Σ≡ refl refl) , Σ≡ refl refl isIdentity = (Σ≡ refl refl) , Σ≡ refl refl
RawFam : RawCategory (lsuc (a b)) (a b) RawFam : RawCategory (lsuc (a b)) (a b)
@ -44,9 +44,9 @@ module _ (a b : Level) where
instance instance
isCategory : IsCategory RawFam isCategory : IsCategory RawFam
isCategory = record isCategory = record
{ assoc = λ {A} {B} {C} {D} {f} {g} {h} assoc {D = D} {f} {g} {h} { isAssociative = λ {A} {B} {C} {D} {f} {g} {h} isAssociative {D = D} {f} {g} {h}
; ident = λ {A} {B} {f} ident {A} {B} {f = f} ; isIdentity = λ {A} {B} {f} isIdentity {A} {B} {f = f}
; arrowIsSet = {!!} ; arrowsAreSets = {!!}
; univalent = {!!} ; univalent = {!!}
} }

View file

@ -9,14 +9,6 @@ open import Cat.Category
open IsCategory open IsCategory
-- 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
-- import Data.List
-- P : (a b : Object ) → Set (')
-- P = {!Data.List.List ?!}
-- Generalized paths:
data Path { ' : Level} {A : Set } (R : A A Set ') : (a b : A) Set ( ') where data Path { ' : Level} {A : Set } (R : A A Set ') : (a b : A) Set ( ') where
empty : {a : A} Path R a a empty : {a : A} Path R a a
cons : {a b c : A} R b c Path R a b Path R a c cons : {a b c : A} R b c Path R a b Path R a c
@ -34,16 +26,16 @@ module _ { ' : Level} ( : Category ') where
open Category open Category
private private
p-assoc : {A B C D : Object} {r : Path Arrow A B} {q : Path Arrow B C} {p : Path Arrow C D} p-isAssociative : {A B C D : Object} {r : Path Arrow A B} {q : Path Arrow B C} {p : Path Arrow C D}
p ++ (q ++ r) (p ++ q) ++ r p ++ (q ++ r) (p ++ q) ++ r
p-assoc {r = r} {q} {empty} = refl p-isAssociative {r = r} {q} {empty} = refl
p-assoc {A} {B} {C} {D} {r = r} {q} {cons x p} = begin p-isAssociative {A} {B} {C} {D} {r = r} {q} {cons x p} = begin
cons x p ++ (q ++ r) ≡⟨ cong (cons x) lem cons x p ++ (q ++ r) ≡⟨ cong (cons x) lem
cons x ((p ++ q) ++ r) ≡⟨⟩ cons x ((p ++ q) ++ r) ≡⟨⟩
(cons x p ++ q) ++ r (cons x p ++ q) ++ r
where where
lem : p ++ (q ++ r) ((p ++ q) ++ r) lem : p ++ (q ++ r) ((p ++ q) ++ r)
lem = p-assoc {r = r} {q} {p} lem = p-isAssociative {r = r} {q} {p}
ident-r : {A} {B} {p : Path Arrow A B} concatenate p empty p ident-r : {A} {B} {p : Path Arrow A B} concatenate p empty p
ident-r {p = empty} = refl ident-r {p = empty} = refl
@ -65,8 +57,8 @@ module _ { ' : Level} ( : Category ') where
} }
RawIsCategoryFree : IsCategory RawFree RawIsCategoryFree : IsCategory RawFree
RawIsCategoryFree = record RawIsCategoryFree = record
{ assoc = λ { {f = f} {g} {h} p-assoc {r = f} {g} {h}} { isAssociative = λ { {f = f} {g} {h} p-isAssociative {r = f} {g} {h}}
; ident = ident-r , ident-l ; isIdentity = ident-r , ident-l
; arrowIsSet = {!!} ; arrowsAreSets = {!!}
; univalent = {!!} ; univalent = {!!}
} }

View file

@ -38,9 +38,6 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
(f : [ A , B ]) (f : [ A , B ])
𝔻 [ θ B F.func→ f ] 𝔻 [ G.func→ f θ A ] 𝔻 [ θ B F.func→ f ] 𝔻 [ G.func→ f θ A ]
-- naturalIsProp : ∀ θ → isProp (Natural θ)
-- naturalIsProp θ x y = {!funExt!}
NaturalTransformation : Set (c c' d') NaturalTransformation : Set (c c' d')
NaturalTransformation = Σ Transformation Natural NaturalTransformation = Σ Transformation Natural
@ -63,8 +60,8 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
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₂ 𝔻.isIdentity
F→ f ≡⟨ sym (proj₁ 𝔻.ident) F→ f ≡⟨ sym (proj₁ 𝔻.isIdentity)
𝔻 [ F→ f 𝟙 𝔻 ] ≡⟨⟩ 𝔻 [ F→ f 𝟙 𝔻 ] ≡⟨⟩
𝔻 [ F→ f identityTrans F A ] 𝔻 [ F→ f identityTrans F A ]
where where
@ -87,11 +84,11 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
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 isAssociative
𝔻 [ θ 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 ] ] ≡⟨ isAssociative
𝔻 [ 𝔻 [ θ 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 isAssociative
𝔻 [ H.func→ f 𝔻 [ θ A η A ] ] ≡⟨⟩ 𝔻 [ H.func→ f 𝔻 [ θ A η A ] ] ≡⟨⟩
𝔻 [ H.func→ f (θ ∘nt η) A ] 𝔻 [ H.func→ f (θ ∘nt η) A ]
where where
@ -100,59 +97,56 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
NatComp = _:⊕:_ NatComp = _:⊕:_
private private
module _ {F G : Functor 𝔻} where module 𝔻 = Category 𝔻
module 𝔻 = Category 𝔻
transformationIsSet : isSet (Transformation F G) module _ {F G : Functor 𝔻} where
transformationIsSet _ _ p q i j C = 𝔻.arrowIsSet _ _ (λ l p l C) (λ l q l C) i j transformationIsSet : isSet (Transformation F G)
IsSet' : { : Level} (A : Set ) Set transformationIsSet _ _ p q i j C = 𝔻.arrowsAreSets _ _ (λ l p l C) (λ l q l C) i j
IsSet' A = {x y : A} (p q : (λ _ A) [ x y ]) p q
naturalIsProp : (θ : Transformation F G) isProp (Natural F G θ) naturalIsProp : (θ : Transformation F G) isProp (Natural F G θ)
naturalIsProp θ θNat θNat' = lem naturalIsProp θ θNat θNat' = lem
where where
lem : (λ _ Natural F G θ) [ (λ f θNat f) (λ f θNat' f) ] lem : (λ _ Natural F G θ) [ (λ f θNat f) (λ f θNat' f) ]
lem = λ i f 𝔻.arrowIsSet _ _ (θNat f) (θNat' f) i lem = λ i f 𝔻.arrowsAreSets _ _ (θNat f) (θNat' f) i
naturalTransformationIsSets : isSet (NaturalTransformation F G) naturalTransformationIsSets : isSet (NaturalTransformation F G)
naturalTransformationIsSets = sigPresSet transformationIsSet naturalTransformationIsSets = sigPresSet transformationIsSet
λ θ ntypeCommulative λ θ ntypeCommulative
(s≤s {n = Nat.suc Nat.zero} z≤n) (s≤s {n = Nat.suc Nat.zero} z≤n)
(naturalIsProp θ) (naturalIsProp θ)
module _ {A B C D : Functor 𝔻} {θ' : NaturalTransformation A B} module _ {A B C D : Functor 𝔻} {θ' : NaturalTransformation A B}
{η' : NaturalTransformation B C} {ζ' : NaturalTransformation C D} where {η' : NaturalTransformation B C} {ζ' : NaturalTransformation C D} where
private private
θ = proj₁ θ' θ = proj₁ θ'
η = proj₁ η' η = proj₁ η'
ζ = proj₁ ζ' ζ = proj₁ ζ'
θNat = proj₂ θ' θNat = proj₂ θ'
ηNat = proj₂ η' ηNat = proj₂ η'
ζNat = proj₂ ζ' ζNat = proj₂ ζ'
L : NaturalTransformation A D L : NaturalTransformation A D
L = (_:⊕:_ {A} {C} {D} ζ' (_:⊕:_ {A} {B} {C} η' θ')) L = (_:⊕:_ {A} {C} {D} ζ' (_:⊕:_ {A} {B} {C} η' θ'))
R : NaturalTransformation A D R : NaturalTransformation A D
R = (_:⊕:_ {A} {B} {D} (_:⊕:_ {B} {C} {D} ζ' η') θ') R = (_:⊕:_ {A} {B} {D} (_:⊕:_ {B} {C} {D} ζ' η') θ')
_g⊕f_ = _:⊕:_ {A} {B} {C} _g⊕f_ = _:⊕:_ {A} {B} {C}
_h⊕g_ = _:⊕:_ {B} {C} {D} _h⊕g_ = _:⊕:_ {B} {C} {D}
:assoc: : L R :isAssociative: : L R
:assoc: = lemSig (naturalIsProp {F = A} {D}) :isAssociative: = lemSig (naturalIsProp {F = A} {D})
L R (funExt (λ x assoc)) L R (funExt (λ x isAssociative))
where where
open Category 𝔻 open Category 𝔻
private
module _ {A B : Functor 𝔻} {f : NaturalTransformation A B} where module _ {A B : Functor 𝔻} {f : NaturalTransformation A B} where
private allNatural = naturalIsProp {F = A} {B}
allNatural = naturalIsProp {F = A} {B} f' = proj₁ f
f' = proj₁ f eq-r : C (𝔻 [ f' C identityTrans A C ]) f' C
module 𝔻Data = Category 𝔻 eq-r C = begin
eq-r : C (𝔻 [ f' C identityTrans A C ]) f' C 𝔻 [ f' C identityTrans A C ] ≡⟨⟩
eq-r C = begin 𝔻 [ f' C 𝔻.𝟙 ] ≡⟨ proj₁ 𝔻.isIdentity
𝔻 [ f' C identityTrans A C ] ≡⟨⟩ f' C
𝔻 [ f' C 𝔻Data.𝟙 ] ≡⟨ proj₁ (𝔻.ident {A} {B}) eq-l : C (𝔻 [ identityTrans B C f' C ]) f' C
f' C eq-l C = proj₂ 𝔻.isIdentity
eq-l : C (𝔻 [ identityTrans B C f' C ]) f' C
eq-l C = proj₂ (𝔻.ident {A} {B})
ident-r : (_:⊕:_ {A} {A} {B} f (identityNat A)) f ident-r : (_:⊕:_ {A} {A} {B} f (identityNat A)) f
ident-r = lemSig allNatural _ _ (funExt eq-r) ident-r = lemSig allNatural _ _ (funExt eq-r)
ident-l : (_:⊕:_ {A} {B} {B} (identityNat B) f) f ident-l : (_:⊕:_ {A} {B} {B} (identityNat B) f) f
@ -174,9 +168,9 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
instance instance
:isCategory: : IsCategory RawFun :isCategory: : IsCategory RawFun
:isCategory: = record :isCategory: = record
{ assoc = λ {A B C D} :assoc: {A} {B} {C} {D} { isAssociative = λ {A B C D} :isAssociative: {A} {B} {C} {D}
; ident = λ {A B} :ident: {A} {B} ; isIdentity = λ {A B} :ident: {A} {B}
; arrowIsSet = λ {F} {G} naturalTransformationIsSets {F} {G} ; arrowsAreSets = λ {F} {G} naturalTransformationIsSets {F} {G}
; univalent = {!!} ; univalent = {!!}
} }

View file

@ -149,10 +149,10 @@ module _ {A B C D : Set} {S : Subset (A × B)} {R : Subset (B × C)} {Q : Subset
(Σ[ b B ] (a , b) S × (Σ[ c C ] (b , c) R × (c , d) Q)) (Σ[ b B ] (a , b) S × (Σ[ c C ] (b , c) R × (c , d) Q))
equi = fwd Cubical.FromStdLib., isequiv equi = fwd Cubical.FromStdLib., isequiv
-- assocc : Q + (R + S) ≡ (Q + R) + S -- isAssociativec : Q + (R + S) ≡ (Q + R) + S
is-assoc : (Σ[ c C ] (Σ[ b B ] (a , b) S × (b , c) R) × (c , d) Q) is-isAssociative : (Σ[ c C ] (Σ[ b B ] (a , b) S × (b , c) R) × (c , d) Q)
(Σ[ b B ] (a , b) S × (Σ[ c C ] (b , c) R × (c , d) Q)) (Σ[ b B ] (a , b) S × (Σ[ c C ] (b , c) R × (c , d) Q))
is-assoc = equivToPath equi is-isAssociative = equivToPath equi
RawRel : RawCategory (lsuc lzero) (lsuc lzero) RawRel : RawCategory (lsuc lzero) (lsuc lzero)
RawRel = record RawRel = record
@ -164,8 +164,8 @@ RawRel = record
RawIsCategoryRel : IsCategory RawRel RawIsCategoryRel : IsCategory RawRel
RawIsCategoryRel = record RawIsCategoryRel = record
{ assoc = funExt is-assoc { isAssociative = funExt is-isAssociative
; ident = funExt ident-l , funExt ident-r ; isIdentity = funExt ident-l , funExt ident-r
; arrowIsSet = {!!} ; arrowsAreSets = {!!}
; univalent = {!!} ; univalent = {!!}
} }

View file

@ -25,10 +25,10 @@ module _ ( : Level) where
_∘_ SetsRaw = Function._∘_ _∘_ SetsRaw = Function._∘_
SetsIsCategory : IsCategory SetsRaw SetsIsCategory : IsCategory SetsRaw
assoc SetsIsCategory = refl isAssociative SetsIsCategory = refl
proj₁ (ident SetsIsCategory) = funExt λ _ refl proj₁ (isIdentity SetsIsCategory) = funExt λ _ refl
proj₂ (ident SetsIsCategory) = funExt λ _ refl proj₂ (isIdentity SetsIsCategory) = funExt λ _ refl
arrowIsSet SetsIsCategory {B = (_ , s)} = setPi λ _ s arrowsAreSets SetsIsCategory {B = (_ , s)} = setPi λ _ s
univalent SetsIsCategory = {!!} univalent SetsIsCategory = {!!}
𝓢𝓮𝓽 Sets : Category (lsuc ) 𝓢𝓮𝓽 Sets : Category (lsuc )
@ -94,12 +94,12 @@ module _ {a b : Level} where
representable : { : Category a b} Category.Object Representable representable : { : Category a b} Category.Object Representable
representable { = } A = record representable { = } A = record
{ raw = record { raw = record
{ func* = λ B [ A , B ] , arrowIsSet { func* = λ B [ A , B ] , arrowsAreSets
; func→ = [_∘_] ; func→ = [_∘_]
} }
; isFunctor = record ; isFunctor = record
{ ident = funExt λ _ proj₂ ident { isIdentity = funExt λ _ proj₂ isIdentity
; distrib = funExt λ x sym assoc ; isDistributive = funExt λ x sym isAssociative
} }
} }
where where
@ -109,12 +109,12 @@ module _ {a b : Level} where
presheaf : { : Category a b} Category.Object (Opposite ) Presheaf presheaf : { : Category a b} Category.Object (Opposite ) Presheaf
presheaf { = } B = record presheaf { = } B = record
{ raw = record { raw = record
{ func* = λ A [ A , B ] , arrowIsSet { func* = λ A [ A , B ] , arrowsAreSets
; func→ = λ f g [ g f ] ; func→ = λ f g [ g f ]
} }
; isFunctor = record ; isFunctor = record
{ ident = funExt λ x proj₁ ident { isIdentity = funExt λ x proj₁ isIdentity
; distrib = funExt λ x assoc ; isDistributive = funExt λ x isAssociative
} }
} }
where where

View file

@ -49,6 +49,9 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
IsIdentity id = {A B : Object} {f : Arrow A B} IsIdentity id = {A B : Object} {f : Arrow A B}
f id f × id f f f id f × id f f
ArrowsAreSets : Set (a b)
ArrowsAreSets = {A B : Object} isSet (Arrow A B)
IsInverseOf : {A B} (Arrow A B) (Arrow B A) Set b IsInverseOf : {A B} (Arrow A B) (Arrow B A) Set b
IsInverseOf = λ f g g f 𝟙 × f g 𝟙 IsInverseOf = λ f g g f 𝟙 × f g 𝟙
@ -80,9 +83,9 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
-- Univalence is indexed by a raw category as well as an identity proof. -- Univalence is indexed by a raw category as well as an identity proof.
module Univalence {a b : Level} ( : RawCategory a b) where module Univalence {a b : Level} ( : RawCategory a b) where
open RawCategory open RawCategory
module _ (ident : IsIdentity 𝟙) where module _ (isIdentity : IsIdentity 𝟙) where
idIso : (A : Object) A A idIso : (A : Object) A A
idIso A = 𝟙 , (𝟙 , ident) idIso A = 𝟙 , (𝟙 , isIdentity)
-- Lemma 9.1.4 in [HoTT] -- Lemma 9.1.4 in [HoTT]
id-to-iso : (A B : Object) A B A B id-to-iso : (A B : Object) A B A B
@ -98,10 +101,10 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
open RawCategory open RawCategory
open Univalence public open Univalence public
field field
assoc : IsAssociative isAssociative : IsAssociative
ident : IsIdentity 𝟙 isIdentity : IsIdentity 𝟙
arrowIsSet : {A B : Object} isSet (Arrow A B) arrowsAreSets : ArrowsAreSets
univalent : Univalent ident univalent : Univalent isIdentity
-- `IsCategory` is a mere proposition. -- `IsCategory` is a mere proposition.
module _ {a b : Level} {C : RawCategory a b} where module _ {a b : Level} {C : RawCategory a b} where
@ -112,12 +115,12 @@ module _ {a b : Level} {C : RawCategory a b} where
open import Cubical.NType.Properties open import Cubical.NType.Properties
propIsAssociative : isProp IsAssociative propIsAssociative : isProp IsAssociative
propIsAssociative x y i = arrowIsSet _ _ x y i propIsAssociative x y i = arrowsAreSets _ _ x y i
propIsIdentity : {f : {A} Arrow A A} isProp (IsIdentity f) propIsIdentity : {f : {A} Arrow A A} isProp (IsIdentity f)
propIsIdentity a b i propIsIdentity a b i
= arrowIsSet _ _ (fst a) (fst b) i = arrowsAreSets _ _ (fst a) (fst b) i
, arrowIsSet _ _ (snd a) (snd b) i , arrowsAreSets _ _ (snd a) (snd b) i
propArrowIsSet : isProp ( {A B} isSet (Arrow A B)) propArrowIsSet : isProp ( {A B} isSet (Arrow A B))
propArrowIsSet a b i = isSetIsProp a b i propArrowIsSet a b i = isSetIsProp a b i
@ -126,9 +129,9 @@ module _ {a b : Level} {C : RawCategory a b} where
propIsInverseOf x y = λ i propIsInverseOf x y = λ i
let let
h : fst x fst y h : fst x fst y
h = arrowIsSet _ _ (fst x) (fst y) h = arrowsAreSets _ _ (fst x) (fst y)
hh : snd x snd y hh : snd x snd y
hh = arrowIsSet _ _ (snd x) (snd y) hh = arrowsAreSets _ _ (snd x) (snd y)
in h i , hh i in h i , hh i
module _ {A B : Object} {f : Arrow A B} where module _ {A B : Object} {f : Arrow A B} where
@ -139,14 +142,14 @@ module _ {a b : Level} {C : RawCategory a b} where
open Cubical.NType.Properties open Cubical.NType.Properties
geq : g g' geq : g g'
geq = begin geq = begin
g ≡⟨ sym (fst ident) g ≡⟨ sym (fst isIdentity)
g 𝟙 ≡⟨ cong (λ φ g φ) (sym ε') g 𝟙 ≡⟨ cong (λ φ g φ) (sym ε')
g (f g') ≡⟨ assoc g (f g') ≡⟨ isAssociative
(g f) g' ≡⟨ cong (λ φ φ g') η (g f) g' ≡⟨ cong (λ φ φ g') η
𝟙 g' ≡⟨ snd ident 𝟙 g' ≡⟨ snd isIdentity
g' g'
propUnivalent : isProp (Univalent ident) propUnivalent : isProp (Univalent isIdentity)
propUnivalent a b i = propPi (λ iso propHasLevel ⟨-2⟩) a b i propUnivalent a b i = propPi (λ iso propHasLevel ⟨-2⟩) a b i
private private
@ -159,23 +162,28 @@ module _ {a b : Level} {C : RawCategory a b} where
-- projections of `IsCategory` - I've arbitrarily chosed to use this -- projections of `IsCategory` - I've arbitrarily chosed to use this
-- result from `x : IsCategory C`. I don't know which (if any) possibly -- result from `x : IsCategory C`. I don't know which (if any) possibly
-- adverse effects this may have. -- adverse effects this may have.
ident : (λ _ IsIdentity 𝟙) [ X.ident Y.ident ] isIdentity : (λ _ IsIdentity 𝟙) [ X.isIdentity Y.isIdentity ]
ident = propIsIdentity x X.ident Y.ident isIdentity = propIsIdentity x X.isIdentity Y.isIdentity
done : x y done : x y
U : {a : IsIdentity 𝟙} (λ _ IsIdentity 𝟙) [ X.ident a ] (b : Univalent a) Set _ U : {a : IsIdentity 𝟙}
U eqwal bbb = (λ i Univalent (eqwal i)) [ X.univalent bbb ] (λ _ IsIdentity 𝟙) [ X.isIdentity a ]
(b : Univalent a)
Set _
U eqwal bbb =
(λ i Univalent (eqwal i))
[ X.univalent bbb ]
P : (y : IsIdentity 𝟙) P : (y : IsIdentity 𝟙)
(λ _ IsIdentity 𝟙) [ X.ident y ] Set _ (λ _ IsIdentity 𝟙) [ X.isIdentity y ] Set _
P y eq = (b' : Univalent y) U eq b' P y eq = (b' : Univalent y) U eq b'
helper : (b' : Univalent X.ident) helper : (b' : Univalent X.isIdentity)
(λ _ Univalent X.ident) [ X.univalent b' ] (λ _ Univalent X.isIdentity) [ X.univalent b' ]
helper univ = propUnivalent x X.univalent univ helper univ = propUnivalent x X.univalent univ
foo = pathJ P helper Y.ident ident foo = pathJ P helper Y.isIdentity isIdentity
eqUni : U ident Y.univalent eqUni : U isIdentity Y.univalent
eqUni = foo Y.univalent eqUni = foo Y.univalent
IC.assoc (done i) = propIsAssociative x X.assoc Y.assoc i IC.isAssociative (done i) = propIsAssociative x X.isAssociative Y.isAssociative i
IC.ident (done i) = ident i IC.isIdentity (done i) = isIdentity i
IC.arrowIsSet (done i) = propArrowIsSet x X.arrowIsSet Y.arrowIsSet i IC.arrowsAreSets (done i) = propArrowIsSet x X.arrowsAreSets Y.arrowsAreSets i
IC.univalent (done i) = eqUni i IC.univalent (done i) = eqUni i
propIsCategory : isProp (IsCategory C) propIsCategory : isProp (IsCategory C)
@ -208,9 +216,9 @@ module _ {a b : Level} ( : Category a b) where
RawCategory._∘_ OpRaw = Function.flip _∘_ RawCategory._∘_ OpRaw = Function.flip _∘_
OpIsCategory : IsCategory OpRaw OpIsCategory : IsCategory OpRaw
IsCategory.assoc OpIsCategory = sym assoc IsCategory.isAssociative OpIsCategory = sym isAssociative
IsCategory.ident OpIsCategory = swap ident IsCategory.isIdentity OpIsCategory = swap isIdentity
IsCategory.arrowIsSet OpIsCategory = arrowIsSet IsCategory.arrowsAreSets OpIsCategory = arrowsAreSets
IsCategory.univalent OpIsCategory = {!!} IsCategory.univalent OpIsCategory = {!!}
Opposite : Category a b Opposite : Category a b
@ -234,9 +242,9 @@ module _ {a b : Level} { : Category a b} where
open IsCategory open IsCategory
module IsCat = IsCategory ( .isCategory) module IsCat = IsCategory ( .isCategory)
rawIsCat : (i : I) IsCategory (rawOp i) rawIsCat : (i : I) IsCategory (rawOp i)
assoc (rawIsCat i) = IsCat.assoc isAssociative (rawIsCat i) = IsCat.isAssociative
ident (rawIsCat i) = IsCat.ident isIdentity (rawIsCat i) = IsCat.isIdentity
arrowIsSet (rawIsCat i) = IsCat.arrowIsSet arrowsAreSets (rawIsCat i) = IsCat.arrowsAreSets
univalent (rawIsCat i) = IsCat.univalent univalent (rawIsCat i) = IsCat.univalent
Opposite-is-involution : Opposite (Opposite ) Opposite-is-involution : Opposite (Opposite )

View file

@ -35,5 +35,6 @@ module _ { '} ( : Category ') {{hasProducts : HasProducts }}
transpose A f = proj₁ (isExponential A f) transpose A f = proj₁ (isExponential A f)
record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where
open Exponential public
field field
exponent : (A B : Object ) Exponential A B exponent : (A B : Object ) Exponential A B

View file

@ -7,7 +7,7 @@ open import Function
open import Cat.Category open import Cat.Category
open Category hiding (_∘_ ; raw) open Category hiding (_∘_ ; raw ; IsIdentity)
module _ {c c' d d'} module _ {c c' d d'}
( : Category c c') ( : Category c c')
@ -23,42 +23,40 @@ module _ {c c' d d'}
func* : Object Object 𝔻 func* : Object Object 𝔻
func→ : {A B} [ A , B ] 𝔻 [ func* A , func* B ] func→ : {A B} [ A , B ] 𝔻 [ func* A , func* B ]
IsIdentity : Set _
IsIdentity = {A : Object } func→ (𝟙 {A}) 𝟙 𝔻 {func* A}
IsDistributive : Set _
IsDistributive = {A B C : Object } {f : [ A , B ]} {g : [ B , C ]}
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ]
record IsFunctor (F : RawFunctor) : 𝓤 where record IsFunctor (F : RawFunctor) : 𝓤 where
open RawFunctor F open RawFunctor F public
field field
ident : {c : Object } func→ (𝟙 {c}) 𝟙 𝔻 {func* c} isIdentity : IsIdentity
distrib : {A B C : Object } {f : [ A , B ]} {g : [ B , C ]} isDistributive : IsDistributive
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ]
record Functor : Set (c c' d d') where record Functor : Set (c c' d d') where
field field
raw : RawFunctor raw : RawFunctor
{{isFunctor}} : IsFunctor raw {{isFunctor}} : IsFunctor raw
private open IsFunctor isFunctor public
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 open Functor
module _ module _
{a b : Level} {a b : Level}
{ 𝔻 : Category a b} { 𝔻 : Category a b}
{F : RawFunctor 𝔻} (F : RawFunctor 𝔻)
where where
private private
module 𝔻 = IsCategory (isCategory 𝔻) module 𝔻 = IsCategory (isCategory 𝔻)
propIsFunctor : isProp (IsFunctor _ _ F) propIsFunctor : isProp (IsFunctor _ _ F)
propIsFunctor isF0 isF1 i = record propIsFunctor isF0 isF1 i = record
{ ident = 𝔻.arrowIsSet _ _ isF0.ident isF1.ident i { isIdentity = 𝔻.arrowsAreSets _ _ isF0.isIdentity isF1.isIdentity i
; distrib = 𝔻.arrowIsSet _ _ isF0.distrib isF1.distrib i ; isDistributive = 𝔻.arrowsAreSets _ _ isF0.isDistributive isF1.isDistributive i
} }
where where
module isF0 = IsFunctor isF0 module isF0 = IsFunctor isF0
@ -77,7 +75,7 @@ module _
IsFunctorIsProp' : IsProp' λ i IsFunctor _ _ (F i) IsFunctorIsProp' : IsProp' λ i IsFunctor _ _ (F i)
IsFunctorIsProp' isF0 isF1 = lemPropF {B = IsFunctor 𝔻} IsFunctorIsProp' isF0 isF1 = lemPropF {B = IsFunctor 𝔻}
(\ F propIsFunctor {F = F}) (\ i F i) (\ F propIsFunctor F) (\ i F i)
where where
open import Cubical.NType.Properties using (lemPropF) open import Cubical.NType.Properties using (lemPropF)
@ -108,8 +106,8 @@ module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : F
dist : (F→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ] dist : (F→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ]
dist = begin dist = begin
(F→ G→) (A [ α1 α0 ]) ≡⟨ refl (F→ G→) (A [ α1 α0 ]) ≡⟨ refl
F→ (G→ (A [ α1 α0 ])) ≡⟨ cong F→ (G .isFunctor .distrib) F→ (G→ (A [ α1 α0 ])) ≡⟨ cong F→ (isDistributive G)
F→ (B [ G→ α1 G→ α0 ]) ≡⟨ F .isFunctor .distrib F→ (B [ G→ α1 G→ α0 ]) ≡⟨ isDistributive F
C [ (F→ G→) α1 (F→ G→) α0 ] C [ (F→ G→) α1 (F→ G→) α0 ]
_∘fr_ : RawFunctor A C _∘fr_ : RawFunctor A C
@ -118,12 +116,12 @@ module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : F
instance instance
isFunctor' : IsFunctor A C _∘fr_ isFunctor' : IsFunctor A C _∘fr_
isFunctor' = record isFunctor' = record
{ ident = begin { isIdentity = begin
(F→ G→) (𝟙 A) ≡⟨ refl (F→ G→) (𝟙 A) ≡⟨ refl
F→ (G→ (𝟙 A)) ≡⟨ cong F→ (G .isFunctor .ident) F→ (G→ (𝟙 A)) ≡⟨ cong F→ (isIdentity G)
F→ (𝟙 B) ≡⟨ F .isFunctor .ident F→ (𝟙 B) ≡⟨ isIdentity F
𝟙 C 𝟙 C
; distrib = dist ; isDistributive = dist
} }
_∘f_ : Functor A C _∘f_ : Functor A C
@ -137,7 +135,7 @@ identity = record
; func→ = λ x x ; func→ = λ x x
} }
; isFunctor = record ; isFunctor = record
{ ident = refl { isIdentity = refl
; distrib = refl ; isDistributive = refl
} }
} }

View file

@ -17,106 +17,77 @@ module _ { ' : Level} { : Category '} { A B : Category.Object
iso-is-epi : Isomorphism f Epimorphism {X = X} f iso-is-epi : Isomorphism f Epimorphism {X = X} f
iso-is-epi (f- , left-inv , right-inv) g₀ g₁ eq = begin iso-is-epi (f- , left-inv , right-inv) g₀ g₁ eq = begin
g₀ ≡⟨ sym (proj₁ ident) g₀ ≡⟨ sym (proj₁ isIdentity)
g₀ 𝟙 ≡⟨ cong (_∘_ g₀) (sym right-inv) g₀ 𝟙 ≡⟨ cong (_∘_ g₀) (sym right-inv)
g₀ (f f-) ≡⟨ assoc g₀ (f f-) ≡⟨ isAssociative
(g₀ f) f- ≡⟨ cong (λ φ φ f-) eq (g₀ f) f- ≡⟨ cong (λ φ φ f-) eq
(g₁ f) f- ≡⟨ sym assoc (g₁ f) f- ≡⟨ sym isAssociative
g₁ (f f-) ≡⟨ cong (_∘_ g₁) right-inv g₁ (f f-) ≡⟨ cong (_∘_ g₁) right-inv
g₁ 𝟙 ≡⟨ proj₁ ident g₁ 𝟙 ≡⟨ proj₁ isIdentity
g₁ g₁
iso-is-mono : Isomorphism f Monomorphism {X = X} f iso-is-mono : Isomorphism f Monomorphism {X = X} f
iso-is-mono (f- , (left-inv , right-inv)) g₀ g₁ eq = iso-is-mono (f- , (left-inv , right-inv)) g₀ g₁ eq =
begin begin
g₀ ≡⟨ sym (proj₂ ident) g₀ ≡⟨ sym (proj₂ isIdentity)
𝟙 g₀ ≡⟨ cong (λ φ φ g₀) (sym left-inv) 𝟙 g₀ ≡⟨ cong (λ φ φ g₀) (sym left-inv)
(f- f) g₀ ≡⟨ sym assoc (f- f) g₀ ≡⟨ sym isAssociative
f- (f g₀) ≡⟨ cong (_∘_ f-) eq f- (f g₀) ≡⟨ cong (_∘_ f-) eq
f- (f g₁) ≡⟨ assoc f- (f g₁) ≡⟨ isAssociative
(f- f) g₁ ≡⟨ cong (λ φ φ g₁) left-inv (f- f) g₁ ≡⟨ cong (λ φ φ g₁) left-inv
𝟙 g₁ ≡⟨ proj₂ ident 𝟙 g₁ ≡⟨ proj₂ isIdentity
g₁ g₁
iso-is-epi-mono : Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f iso-is-epi-mono : Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f
iso-is-epi-mono iso = iso-is-epi iso , iso-is-mono iso iso-is-epi-mono iso = iso-is-epi iso , iso-is-mono iso
{- -- TODO: We want to avoid defining the yoneda embedding going through the
epi-mono-is-not-iso : { '} ¬ (( : Category {} {'}) {A B X : Object } (f : Arrow A B ) Epimorphism { = } {X = X} f Monomorphism { = } {X = X} f Isomorphism { = } f) -- category of categories (since it doesn't exist).
epi-mono-is-not-iso f = open import Cat.Categories.Cat using (RawCat)
let k = f {!!} {!!} {!!} {!!}
in {!!}
-}
open import Cat.Category module _ { : Level} { : Category } (unprovable : IsCategory (RawCat )) where
open Category open import Cat.Categories.Fun
open Functor open import Cat.Categories.Sets
module Cat = Cat.Categories.Cat
open import Cat.Category.Exponential
open Functor
𝓢 = Sets
private
Cat : Category _ _
Cat = record { raw = RawCat ; isCategory = unprovable}
prshf = presheaf { = }
module = Category
-- module _ { : Level} { : Category } _⇑_ : (A B : Category.Object Cat) Category.Object Cat
-- {isSObj : isSet ( .Object)} A B = (exponent A B) .obj
-- {isz2 : ∀ {} → {A B : Set } → isSet (Sets [ A , B ])} where where
-- -- open import Cat.Categories.Cat using (Cat) open HasExponentials (Cat.hasExponentials unprovable)
-- open import Cat.Categories.Fun
-- open import Cat.Categories.Sets
-- -- module Cat = Cat.Categories.Cat
-- open import Cat.Category.Exponential
-- private
-- Cat = Cat
-- prshf = presheaf { = }
-- module = IsCategory ( .isCategory)
-- -- Exp : Set (lsuc (lsuc )) module _ {A B : .Object} (f : [ A , B ]) where
-- -- Exp = Exponential (Cat (lsuc ) ) :func→: : NaturalTransformation (prshf A) (prshf B)
-- -- Sets (Opposite ) :func→: = (λ C x [ f x ]) , λ f₁ funExt λ _ .isAssociative
-- _⇑_ : (A B : Cat .Object) → Cat .Object module _ {c : Category.Object } where
-- A ⇑ B = (exponent A B) .obj eqTrans : (λ _ Transformation (prshf c) (prshf c))
-- where [ (λ _ x [ .𝟙 x ]) identityTrans (prshf c) ]
-- open HasExponentials (Cat.hasExponentials ) eqTrans = funExt λ x funExt λ x .isIdentity .proj₂
-- module _ {A B : .Object} (f : .Arrow A B) where open import Cubical.NType.Properties
-- :func→: : NaturalTransformation (prshf A) (prshf B) open import Cat.Categories.Fun
-- :func→: = (λ C x [ f x ]) , λ f₁ funExt λ _ .assoc :ident: : :func→: (.𝟙 {c}) Category.𝟙 Fun {A = prshf c}
:ident: = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq
where
eq : (λ C x [ .𝟙 x ]) identityTrans (prshf c)
eq = funExt λ A funExt λ B proj₂ .isIdentity
-- module _ {c : .Object} where yoneda : Functor (Fun { = Opposite } {𝔻 = 𝓢})
-- eqTrans : (λ _ → Transformation (prshf c) (prshf c)) yoneda = record
-- [ (λ _ x → [ .𝟙 ∘ x ]) ≡ identityTrans (prshf c) ] { raw = record
-- eqTrans = funExt λ x → funExt λ x → .ident .proj₂ { func* = prshf
; func→ = :func→:
-- eqNat : (λ i → Natural (prshf c) (prshf c) (eqTrans i)) }
-- [(λ _ → funExt (λ _ → .assoc)) ≡ identityNatural (prshf c)] ; isFunctor = record
-- eqNat = λ i {A} {B} f → { isIdentity = :ident:
-- let ; isDistributive = {!!}
-- open IsCategory (Sets .isCategory) }
-- lemm : (Sets [ eqTrans i B ∘ prshf c .func→ f ]) ≡ }
-- (Sets [ prshf c .func→ f ∘ eqTrans i A ])
-- lemm = {!!}
-- lem : (λ _ → Sets [ Functor.func* (prshf c) A , prshf c .func* B ])
-- [ Sets [ eqTrans i B ∘ prshf c .func→ f ]
-- ≡ Sets [ prshf c .func→ f ∘ eqTrans i A ] ]
-- lem
-- = isz2 _ _ lemm _ i
-- -- (Sets [ eqTrans i B ∘ prshf c .func→ f ])
-- -- (Sets [ prshf c .func→ f ∘ eqTrans i A ])
-- -- lemm
-- -- _ i
-- in
-- lem
-- -- eqNat = λ {A} {B} i [B,A] i' [A,c] →
-- -- let
-- -- k : [ {!!} , {!!} ]
-- -- k = [A,c]
-- -- in {! [ ? ∘ ? ]!}
-- :ident: : (:func→: ( .𝟙 {c})) (Fun .𝟙 {o = prshf c})
-- :ident: = Σ≡ eqTrans eqNat
-- yoneda : Functor (Fun { = Opposite } {𝔻 = Sets {}})
-- yoneda = record
-- { func* = prshf
-- ; func→ = :func→:
-- ; isFunctor = record
-- { ident = :ident:
-- ; distrib = {!!}
-- }
-- }

View file

@ -20,28 +20,3 @@ module Equality where
Σ≡ : a b Σ≡ : a b
proj₁ (Σ≡ i) = proj₁≡ i proj₁ (Σ≡ i) = proj₁≡ i
proj₂ (Σ≡ i) = proj₂≡ i proj₂ (Σ≡ i) = proj₂≡ i
-- Remark 2.7.1: This theorem:
--
-- (x , u) ≡ (x , v) → u ≡ v
--
-- does *not* hold! We can only conclude that there *exists* `p : x ≡ x`
-- such that
--
-- p* u ≡ v
-- thm : isSet A → (∀ {a} → isSet (B a)) → isSet (Σ A B)
-- thm sA sB (x , y) (x' , y') p q = res
-- where
-- x≡x'0 : x ≡ x'
-- x≡x'0 = λ i → proj₁ (p i)
-- x≡x'1 : x ≡ x'
-- x≡x'1 = λ i → proj₁ (q i)
-- someP : x ≡ x'
-- someP = {!!}
-- tricky : {!y!} ≡ y'
-- tricky = {!!}
-- -- res' : (λ _ → Σ A B) [ (x , y) ≡ (x' , y') ]
-- res' : ({!!} , {!!}) ≡ ({!!} , {!!})
-- res' = {!!}
-- res : p ≡ q
-- res i = {!res'!}