Refactor category of categories

No longer actually define the category. Just define the raw category and
a few results about it.
This commit is contained in:
Frederik Hanghøj Iversen 2018-02-23 10:34:37 +01:00
parent 32b9ce2ea8
commit a87d404aad

View file

@ -86,34 +86,43 @@ module _ ( ' : Level) where
ident-l : identity ∘f F F ident-l : identity ∘f F F
ident-l = Functor≡ eq* eq→ ident-l = Functor≡ eq* eq→
RawCat : RawCategory (lsuc ( ')) ( ') RawCat : RawCategory (lsuc ( ')) ( ')
RawCat = 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 private
instance open RawCategory
:isCategory: : IsCategory RawCat assoc : IsAssociative RawCat
assoc :isCategory: {f = F} {G} {H} = assc {F = F} {G = G} {H = H} assoc {f = F} {G} {H} = assc {F = F} {G = G} {H = H}
ident :isCategory: = ident-r , ident-l -- TODO: Rename `ident'` to `ident` after changing how names are exposed in Functor.
arrowIsSet :isCategory: = {!!} ident' : IsIdentity RawCat identity
univalent :isCategory: = {!!} 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.
Cat : Category (lsuc ( ')) ( ') -- Because of the note above there is not category of categories.
Category.raw Cat = RawCat Cat : (unprovable : IsCategory RawCat) Category (lsuc ( ')) ( ')
Category.raw (Cat _) = RawCat
Category.isCategory (Cat unprovable) = unprovable
-- Category.raw Cat _ = RawCat
-- Category.isCategory Cat unprovable = unprovable
module _ { ' : Level} where -- 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') = [ c , c' ] × 𝔻 [ d , d' ] :Arrow: (c , d) (c' , d') = [ c , c' ] × 𝔻 [ d , d' ]
@ -131,27 +140,23 @@ module _ { ' : Level} where
RawCategory.Arrow :rawProduct: = :Arrow: RawCategory.Arrow :rawProduct: = :Arrow:
RawCategory.𝟙 :rawProduct: = :𝟙: RawCategory.𝟙 :rawProduct: = :𝟙:
RawCategory._∘_ :rawProduct: = _:⊕:_ RawCategory._∘_ :rawProduct: = _:⊕:_
open RawCategory :rawProduct:
module C = Category module C = Category
module D = Category 𝔻 module D = Category 𝔻
postulate postulate
issSet : {A B : RawCategory.Object :rawProduct:} isSet (RawCategory.Arrow :rawProduct: A B) issSet : {A B : RawCategory.Object :rawProduct:} isSet (Arrow A B)
ident' : IsIdentity :𝟙:
ident'
= Σ≡ (fst C.ident) (fst D.ident)
, Σ≡ (snd C.ident) (snd D.ident)
postulate univalent : Univalence.Univalent :rawProduct: ident'
instance instance
:isCategory: : IsCategory :rawProduct: :isCategory: : IsCategory :rawProduct:
-- :isCategory: = record
-- { assoc = Σ≡ C.assoc D.assoc
-- ; ident
-- = Σ≡ (fst C.ident) (fst D.ident)
-- , Σ≡ (snd C.ident) (snd D.ident)
-- ; arrow-is-set = issSet
-- ; univalent = {!!}
-- }
IsCategory.assoc :isCategory: = Σ≡ C.assoc D.assoc IsCategory.assoc :isCategory: = Σ≡ C.assoc D.assoc
IsCategory.ident :isCategory: IsCategory.ident :isCategory: = ident'
= Σ≡ (fst C.ident) (fst D.ident)
, Σ≡ (snd C.ident) (snd D.ident)
IsCategory.arrowIsSet :isCategory: = issSet IsCategory.arrowIsSet :isCategory: = issSet
IsCategory.univalent :isCategory: = {!!} IsCategory.univalent :isCategory: = univalent
:product: : Category ' :product: : Category '
Category.raw :product: = :rawProduct: Category.raw :product: = :rawProduct:
@ -209,32 +214,33 @@ 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 𝔻
@ -276,7 +282,7 @@ module _ ( : Level) where
result : 𝔻 [ func* F A , func* G 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
@ -303,109 +309,109 @@ module _ ( : Level) where
open module 𝔻 = Category 𝔻 open module 𝔻 = Category 𝔻
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₁
B = G×B .proj₂ -- B = G×B .proj₂
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⊕_ = Category._∘_ (Product.obj (:obj: ×p )) {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 × [ A , B ]} -- {θ×f : NaturalTransformation F G × [ A , B ]}
{η×g : NaturalTransformation G H × [ 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 : [ 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 : [ B , C ] -- g : [ B , C ]
g = proj₂ η×g -- g = proj₂ η×g
ηθNT : NaturalTransformation F H -- ηθNT : NaturalTransformation F H
ηθNT = Category._∘_ Fun {F} {G} {H} (η , ηNat) (θ , θNat) -- ηθNT = Category._∘_ Fun {F} {G} {H} (η , ηNat) (θ , θNat)
ηθ = proj₁ ηθNT -- ηθ = proj₁ ηθNT
ηθNat = proj₂ ηθNT -- ηθNat = proj₂ ηθNT
:distrib: : -- :distrib: :
𝔻 [ 𝔻 [ η C θ C ] func→ F ( [ g f ] ) ] -- 𝔻 [ 𝔻 [ η C ∘ θ C ] ∘ func→ F ( [ g ∘ f ] ) ]
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ θ B func→ F f ] ] -- ≡ 𝔻 [ 𝔻 [ η C ∘ func→ G g ] ∘ 𝔻 [ θ B ∘ func→ F f ] ]
:distrib: = begin -- :distrib: = begin
𝔻 [ (ηθ C) func→ F ( [ g f ]) ] -- 𝔻 [ (ηθ C) ∘ func→ F ( [ g ∘ f ]) ]
≡⟨ ηθNat ( [ g f ]) -- ≡⟨ ηθNat ( [ g ∘ f ]) ⟩
𝔻 [ func→ H ( [ g f ]) (ηθ A) ] -- 𝔻 [ func→ H ( [ g ∘ f ]) ∘ (ηθ A) ]
≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.distrib) -- ≡⟨ cong (λ φ → 𝔻 [ φ ∘ ηθ A ]) (H.distrib) ⟩
𝔻 [ 𝔻 [ func→ H g func→ H f ] (ηθ A) ] -- 𝔻 [ 𝔻 [ func→ H g ∘ func→ H f ] ∘ (ηθ A) ]
≡⟨ sym assoc -- ≡⟨ sym assoc ⟩
𝔻 [ func→ H g 𝔻 [ func→ H f ηθ A ] ] -- 𝔻 [ func→ H g ∘ 𝔻 [ func→ H f ∘ ηθ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) assoc -- ≡⟨ cong (λ φ → 𝔻 [ func→ H g ∘ φ ]) assoc ⟩
𝔻 [ func→ H g 𝔻 [ 𝔻 [ func→ H f η A ] θ A ] ] -- 𝔻 [ func→ H g ∘ 𝔻 [ 𝔻 [ func→ H f ∘ η A ] ∘ θ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f))) -- ≡⟨ cong (λ φ → 𝔻 [ func→ H g ∘ φ ]) (cong (λ φ → 𝔻 [ φ ∘ θ A ]) (sym (ηNat f))) ⟩
𝔻 [ func→ H g 𝔻 [ 𝔻 [ η B func→ G f ] θ A ] ] -- 𝔻 [ func→ H g ∘ 𝔻 [ 𝔻 [ η B ∘ func→ G f ] ∘ θ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) (sym assoc) -- ≡⟨ cong (λ φ → 𝔻 [ func→ H g ∘ φ ]) (sym assoc) ⟩
𝔻 [ func→ H g 𝔻 [ η B 𝔻 [ func→ G f θ A ] ] ] -- 𝔻 [ func→ H g ∘ 𝔻 [ η B ∘ 𝔻 [ func→ G f ∘ θ A ] ] ]
≡⟨ assoc -- ≡⟨ assoc ⟩
𝔻 [ 𝔻 [ func→ H g η B ] 𝔻 [ func→ G f θ A ] ] -- 𝔻 [ 𝔻 [ func→ H g ∘ η B ] ∘ 𝔻 [ func→ G f ∘ θ A ] ]
≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ func→ G f θ A ] ]) (sym (ηNat g)) -- ≡⟨ cong (λ φ → 𝔻 [ φ ∘ 𝔻 [ func→ G f ∘ θ A ] ]) (sym (ηNat g)) ⟩
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ func→ G f θ A ] ] -- 𝔻 [ 𝔻 [ η C ∘ func→ G g ] ∘ 𝔻 [ func→ G f ∘ θ A ] ]
≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C func→ G g ] φ ]) (sym (θNat f)) -- ≡⟨ cong (λ φ → 𝔻 [ 𝔻 [ η C ∘ func→ G g ] ∘ φ ]) (sym (θNat f)) ⟩
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ θ B func→ F f ] ] -- 𝔻 [ 𝔻 [ η C ∘ func→ G g ] ∘ 𝔻 [ θ B ∘ func→ F f ] ] ∎
where -- where
open Category 𝔻 -- open Category 𝔻
module H = IsFunctor (H .isFunctor) -- module H = IsFunctor (H .isFunctor)
:eval: : Functor ((:obj: ×p ) .Product.obj) 𝔻 -- :eval: : Functor ((:obj: ×p ) .Product.obj) 𝔻
:eval: = record -- :eval: = record
{ raw = record -- { raw = record
{ func* = :func*: -- { func* = :func*:
; func→ = λ {dom} {cod} :func→: {dom} {cod} -- ; func→ = λ {dom} {cod} → :func→: {dom} {cod}
} -- }
; isFunctor = record -- ; isFunctor = record
{ ident = λ {o} :ident: {o} -- { ident = λ {o} → :ident: {o}
; distrib = λ {f u n k y} :distrib: {f} {u} {n} {k} {y} -- ; distrib = λ {f u n k y} → :distrib: {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 {} {}) renaming (_|×|_ to parallelProduct)
postulate -- postulate
transpose : Functor 𝔸 :obj: -- transpose : Functor 𝔸 :obj:
eq : Cat [ :eval: (parallelProduct transpose (𝟙 Cat {A = })) ] 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)
-- (𝟙 Cat) -- -- (𝟙 Cat)
-- ]) -- -- ])
-- ≡ F -- -- ≡ F
-- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758` -- -- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758`
-- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Cat [ -- -- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Cat [
-- :eval: (parallelProduct F~ (𝟙 Cat {o = }))] F) catTranspose = -- -- :eval: ∘ (parallelProduct F~ (𝟙 Cat {o = }))] ≡ F) catTranspose =
-- transpose , eq -- -- transpose , eq
:isExponential: : IsExponential Cat 𝔻 :obj: :eval: -- :isExponential: : IsExponential Cat 𝔻 :obj: :eval:
:isExponential: = {!catTranspose!} -- :isExponential: = {!catTranspose!}
where -- where
open HasProducts (hasProducts {} {}) using (_|×|_) -- open HasProducts (hasProducts {} {}) using (_|×|_)
-- :isExponential: = λ 𝔸 F transpose 𝔸 F , eq' 𝔸 F -- -- :isExponential: = λ 𝔸 F → transpose 𝔸 F , eq' 𝔸 F
-- :exponent: : Exponential (Cat ) A B -- -- :exponent: : Exponential (Cat ) A B
:exponent: : Exponential Cat 𝔻 -- :exponent: : Exponential Cat 𝔻
:exponent: = record -- :exponent: = record
{ obj = :obj: -- { obj = :obj:
; eval = :eval: -- ; eval = :eval:
; isExponential = :isExponential: -- ; isExponential = :isExponential:
} -- }
hasExponentials : HasExponentials (Cat ) -- hasExponentials : HasExponentials (Cat )
hasExponentials = record { exponent = :exponent: } -- hasExponentials = record { exponent = :exponent: }