cat/src/Cat/Categories/Cat.agda

379 lines
14 KiB
Agda
Raw Normal View History

2018-02-02 14:33:54 +00:00
-- There is no category of categories in our interpretation
{-# OPTIONS --cubical --allow-unsolved-metas #-}
module Cat.Categories.Cat where
open import Agda.Primitive
open import Cubical
open import Function
open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
open import Cat.Category
2018-02-05 15:35:33 +00:00
open import Cat.Category.Functor
open import Cat.Category.Product
open import Cat.Category.Exponential hiding (_×_ ; product)
open import Cat.Category.NaturalTransformation
2018-01-31 13:39:54 +00:00
open import Cat.Equality
open Equality.Data.Product
2018-02-23 11:41:15 +00:00
open Functor using (func→ ; func*)
2018-02-22 14:31:54 +00:00
open Category using (Object ; 𝟙)
2018-01-25 11:01:37 +00:00
-- The category of categories
2018-01-24 15:38:28 +00:00
module _ ( ' : Level) where
private
2018-01-31 13:39:54 +00:00
module _ {𝔸 𝔹 𝔻 : Category '} {F : Functor 𝔸 𝔹} {G : Functor 𝔹 } {H : Functor 𝔻} where
assc : F[ H F[ G F ] ] F[ F[ H G ] F ]
2018-03-05 16:10:41 +00:00
assc = Functor≡ refl
2018-01-25 11:44:47 +00:00
module _ { 𝔻 : Category '} {F : Functor 𝔻} where
ident-r : F[ F identity ] F
2018-03-05 16:10:41 +00:00
ident-r = Functor≡ refl
2018-02-23 11:15:16 +00:00
ident-l : F[ identity F ] F
2018-03-05 16:10:41 +00:00
ident-l = Functor≡ refl
RawCat : RawCategory (lsuc ( ')) ( ')
RawCat =
record
{ Object = Category '
; Arrow = Functor
; 𝟙 = identity
; _∘_ = F[_∘_]
}
private
2018-02-23 11:15:16 +00:00
open RawCategory RawCat
2018-02-23 11:43:49 +00:00
isAssociative : IsAssociative
isAssociative {f = F} {G} {H} = assc {F = F} {G = G} {H = H}
2018-03-08 00:09:40 +00:00
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.
-- 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 CatProduct { ' : Level} ( 𝔻 : Category ') where
private
:Object: = Object × Object 𝔻
:Arrow: : :Object: :Object: Set '
:Arrow: (c , d) (c' , d') = [ c , c' ] × 𝔻 [ d , d' ]
:𝟙: : {o : :Object:} :Arrow: o o
:𝟙: = 𝟙 , 𝟙 𝔻
_:⊕:_ :
{a b c : :Object:}
:Arrow: b c
:Arrow: a b
:Arrow: a c
_:⊕:_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) [ bc∈C ab∈C ] , 𝔻 [ bc∈D ab∈D ]}
:rawProduct: : RawCategory '
RawCategory.Object :rawProduct: = :Object:
RawCategory.Arrow :rawProduct: = :Arrow:
RawCategory.𝟙 :rawProduct: = :𝟙:
RawCategory._∘_ :rawProduct: = _:⊕:_
open RawCategory :rawProduct:
module = Category
module 𝔻 = Category 𝔻
open import Cubical.Sigma
arrowsAreSets : ArrowsAreSets -- {A B : RawCategory.Object :rawProduct:} → isSet (Arrow A B)
arrowsAreSets = setSig {sA = .arrowsAreSets} {sB = λ x 𝔻.arrowsAreSets}
isIdentity : IsIdentity :𝟙:
isIdentity
= Σ≡ (fst .isIdentity) (fst 𝔻.isIdentity)
, Σ≡ (snd .isIdentity) (snd 𝔻.isIdentity)
postulate univalent : Univalence.Univalent :rawProduct: isIdentity
instance
:isCategory: : IsCategory :rawProduct:
IsCategory.isAssociative :isCategory: = Σ≡ .isAssociative 𝔻.isAssociative
IsCategory.isIdentity :isCategory: = isIdentity
IsCategory.arrowsAreSets :isCategory: = arrowsAreSets
IsCategory.univalent :isCategory: = univalent
obj : Category '
Category.raw obj = :rawProduct:
proj₁ : Functor obj
proj₁ = record
{ raw = record { func* = fst ; func→ = fst }
; isFunctor = record { isIdentity = refl ; isDistributive = refl }
}
proj₂ : Functor obj 𝔻
proj₂ = record
{ raw = record { func* = snd ; func→ = snd }
; isFunctor = record { isIdentity = refl ; isDistributive = refl }
}
module _ {X : Category '} (x₁ : Functor X ) (x₂ : Functor X 𝔻) where
2018-01-24 15:38:28 +00:00
private
x : Functor X obj
x = record
{ raw = record
{ func* = λ x x₁.func* x , x₂.func* x
; func→ = λ x x₁.func→ x , x₂.func→ x
}
; isFunctor = record
{ isIdentity = Σ≡ x₁.isIdentity x₂.isIdentity
; isDistributive = Σ≡ x₁.isDistributive x₂.isDistributive
}
}
where
open module x = Functor x₁
open module x = Functor x₂
isUniqL : F[ proj₁ x ] x₁
2018-03-05 16:10:41 +00:00
isUniqL = Functor≡ refl
2018-01-24 15:38:28 +00:00
isUniqR : F[ proj₂ x ] x₂
2018-03-05 16:10:41 +00:00
isUniqR = Functor≡ refl
2018-01-24 15:38:28 +00:00
isUniq : F[ proj₁ x ] x₁ × F[ proj₂ x ] x₂
isUniq = isUniqL , isUniqR
2018-01-24 15:38:28 +00:00
isProduct : ∃![ x ] (F[ proj₁ x ] x₁ × F[ proj₂ x ] x₂)
isProduct = x , isUniq
2018-01-24 15:38:28 +00:00
module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
private
Cat = Cat ' unprovable
module _ ( 𝔻 : Category ') where
private
module P = CatProduct 𝔻
2018-03-08 09:22:21 +00:00
rawProduct : RawProduct Cat 𝔻
2018-03-08 09:20:29 +00:00
RawProduct.obj rawProduct = P.obj
RawProduct.proj₁ rawProduct = P.proj₁
RawProduct.proj₂ rawProduct = P.proj₂
isProduct : IsProduct Cat rawProduct
IsProduct.isProduct isProduct = P.isProduct
2018-01-24 15:38:28 +00:00
2018-03-08 09:22:21 +00:00
product : Product Cat 𝔻
2018-03-08 09:20:29 +00:00
Product.raw product = rawProduct
Product.isProduct product = isProduct
2018-01-24 15:38:28 +00:00
instance
hasProducts : HasProducts Cat
hasProducts = record { product = product }
2018-01-24 15:38:28 +00:00
-- Basically proves that `Cat ` is cartesian closed.
module CatExponential { : Level} ( 𝔻 : Category ) where
open Data.Product
open import Cat.Categories.Fun
Category = Category
open Fun 𝔻 renaming (identity to idN)
2018-01-24 15:38:28 +00:00
private
:func*: : Functor 𝔻 × Object Object 𝔻
:func*: (F , A) = func* F A
2018-01-25 11:01:37 +00:00
prodObj : Category
prodObj = Fun
module _ {dom cod : Functor 𝔻 × Object } where
private
F : Functor 𝔻
F = proj₁ dom
A : Object
A = proj₂ dom
G : Functor 𝔻
G = proj₁ cod
B : Object
B = proj₂ cod
:func→: : (pobj : NaturalTransformation F G × [ A , B ])
𝔻 [ func* F A , func* G B ]
:func→: ((θ , θNat) , f) = result
where
θA : 𝔻 [ func* F A , func* G A ]
θA = θ A
θB : 𝔻 [ func* F B , func* G B ]
θB = θ B
F→f : 𝔻 [ func* F A , func* F B ]
F→f = func→ F f
G→f : 𝔻 [ func* G A , func* G B ]
G→f = func→ G f
l : 𝔻 [ func* F A , func* G B ]
l = 𝔻 [ θB F→f ]
r : 𝔻 [ func* F A , func* G B ]
r = 𝔻 [ G→f θA ]
-- There are two choices at this point,
-- but I suppose the whole point is that
-- by `θNat f` we have `l ≡ r`
-- lem : 𝔻 [ θ B ∘ F .func→ f ] ≡ 𝔻 [ G .func→ f ∘ θ A ]
-- lem = θNat f
result : 𝔻 [ func* F A , func* G B ]
result = l
open CatProduct renaming (obj to _×p_) using ()
module _ {c : Functor 𝔻 × Object } where
private
F : Functor 𝔻
F = proj₁ c
C : Object
C = proj₂ c
-- NaturalTransformation F G × .Arrow A B
-- :ident: : :func→: {c} {c} (identityNat F , .𝟙) 𝔻 .𝟙
-- :ident: = trans (proj₂ 𝔻.isIdentity) (F .isIdentity)
-- where
-- open module 𝔻 = IsCategory (𝔻 .isCategory)
-- Unfortunately the equational version has some ambigous arguments.
:ident: : :func→: {c} {c} (NT.identity F , 𝟙 {A = proj₂ c}) 𝟙 𝔻
:ident: = begin
:func→: {c} {c} (𝟙 (prodObj ×p ) {c}) ≡⟨⟩
:func→: {c} {c} (idN F , 𝟙 ) ≡⟨⟩
𝔻 [ identityTrans F C func→ F (𝟙 )] ≡⟨⟩
𝔻 [ 𝟙 𝔻 func→ F (𝟙 )] ≡⟨ proj₂ 𝔻.isIdentity
func→ F (𝟙 ) ≡⟨ F.isIdentity
𝟙 𝔻
where
open module 𝔻 = Category 𝔻
open module F = Functor F
module _ {F×A G×B H×C : Functor 𝔻 × Object } where
F = F×A .proj₁
A = F×A .proj₂
G = G×B .proj₁
B = G×B .proj₂
H = H×C .proj₁
C = H×C .proj₂
-- Not entirely clear what this is at this point:
_P⊕_ = Category._∘_ (prodObj ×p ) {F×A} {G×B} {H×C}
module _
-- NaturalTransformation F G × .Arrow A B
{θ×f : NaturalTransformation F G × [ A , B ]}
{η×g : NaturalTransformation G H × [ B , C ]} where
2018-01-24 15:38:28 +00:00
private
θ : Transformation F G
θ = proj₁ (proj₁ θ×f)
θNat : Natural F G θ
θNat = proj₂ (proj₁ θ×f)
f : [ A , B ]
f = proj₂ θ×f
η : Transformation G H
η = proj₁ (proj₁ η×g)
ηNat : Natural G H η
ηNat = proj₂ (proj₁ η×g)
g : [ B , C ]
g = proj₂ η×g
ηθNT : NaturalTransformation F H
ηθNT = Category._∘_ Fun {F} {G} {H} (η , ηNat) (θ , θNat)
ηθ = proj₁ ηθNT
ηθNat = proj₂ ηθNT
:isDistributive: :
𝔻 [ 𝔻 [ η C θ C ] func→ F ( [ g f ] ) ]
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ θ B func→ F f ] ]
:isDistributive: = begin
𝔻 [ (ηθ C) func→ F ( [ g f ]) ]
≡⟨ ηθNat ( [ g f ])
𝔻 [ func→ H ( [ g f ]) (ηθ A) ]
≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.isDistributive)
𝔻 [ 𝔻 [ func→ H g func→ H f ] (ηθ A) ]
≡⟨ sym isAssociative
𝔻 [ func→ H g 𝔻 [ func→ H f ηθ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) isAssociative
𝔻 [ func→ H g 𝔻 [ 𝔻 [ func→ H f η A ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f)))
𝔻 [ func→ H g 𝔻 [ 𝔻 [ η B func→ G f ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) (sym isAssociative)
𝔻 [ func→ H g 𝔻 [ η B 𝔻 [ func→ G f θ A ] ] ]
≡⟨ isAssociative
𝔻 [ 𝔻 [ func→ H g η B ] 𝔻 [ func→ G f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ func→ G f θ A ] ]) (sym (ηNat g))
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ func→ G f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C func→ G g ] φ ]) (sym (θNat f))
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ θ B func→ F f ] ]
where
open Category 𝔻
module H = Functor H
eval : Functor (CatProduct.obj prodObj ) 𝔻
-- :eval: : Functor (prodObj ×p ) 𝔻
eval = record
{ raw = record
{ func* = :func*:
; func→ = λ {dom} {cod} :func→: {dom} {cod}
}
; isFunctor = record
{ isIdentity = λ {o} :ident: {o}
; isDistributive = λ {f u n k y} :isDistributive: {f} {u} {n} {k} {y}
}
}
2018-02-23 09:44:23 +00:00
module _ (𝔸 : Category ) (F : Functor (𝔸 ×p ) 𝔻) where
-- open HasProducts (hasProducts {} {} unprovable) renaming (_|×|_ to parallelProduct)
postulate
parallelProduct
: Functor 𝔸 prodObj Functor
Functor (𝔸 ×p ) (prodObj ×p )
transpose : Functor 𝔸 prodObj
eq : F[ eval (parallelProduct transpose (identity {C = })) ] F
-- eq : F[ :eval: ∘ {!!} ] ≡ F
-- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (𝟙 Cat {o = })) ] ≡ F
-- eq' : (Cat [ :eval: ∘
-- (record { product = product } HasProducts.|×| transpose)
-- (𝟙 Cat)
-- ])
-- ≡ F
-- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758`
-- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Cat [
-- :eval: (parallelProduct F~ (𝟙 Cat {o = }))] F) catTranspose =
-- transpose , eq
module _ ( : Level) (unprovable : IsCategory (RawCat )) where
private
Cat : Category (lsuc ( )) ( )
Cat = Cat unprovable
module _ ( 𝔻 : Category ) where
open CatExponential 𝔻 using (prodObj ; eval)
-- Putting in the type annotation causes Agda to loop indefinitely.
-- eval' : Functor (CatProduct.obj prodObj ) 𝔻
-- Likewise, using it below also results in this.
eval' : _
eval' = eval
-- private
-- -- module _ ( 𝔻 : Category ) where
-- postulate :isExponential: : IsExponential Cat 𝔻 prodObj :eval:
-- -- :isExponential: : IsExponential Cat 𝔻 :obj: :eval:
-- -- :isExponential: = {!catTranspose!}
-- -- where
-- -- open HasProducts (hasProducts {} {} unprovable) using (_|×|_)
-- -- :isExponential: = λ 𝔸 F → transpose 𝔸 F , eq' 𝔸 F
-- -- :exponent: : Exponential (Cat ) A B
exponent : Exponential Cat 𝔻
exponent = record
{ obj = prodObj
; eval = {!evalll'!}
; isExponential = {!:isExponential:!}
}
where
open HasProducts (hasProducts unprovable) renaming (_×_ to _×p_)
open import Cat.Categories.Fun
open Fun
-- _×p_ = CatProduct.obj -- prodObj
-- eval' : Functor CatP.obj 𝔻
2018-02-23 09:44:23 +00:00
hasExponentials : HasExponentials Cat
hasExponentials = record { exponent = exponent }