cat/src/Cat/Categories/Cat.agda

371 lines
13 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 Data.Product renaming (proj₁ to fst ; proj₂ to snd)
2018-03-08 10:54:13 +00:00
open import Cubical
open import Cubical.Sigma
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-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
2018-03-08 10:54:13 +00:00
-- 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 this there is no category of categories.
Cat : (unprovable : IsCategory RawCat) Category (lsuc ( ')) ( ')
Category.raw (Cat _) = RawCat
Category.isCategory (Cat unprovable) = unprovable
2018-03-08 10:54:13 +00:00
-- | In the following we will pretend there is a category of categories when
-- e.g. talking about it being cartesian closed. It still makes sense to
-- construct these things even though that category does not exist.
--
-- If the notion of a category is later generalized to work on different
-- homotopy levels, then the proof that the category of categories is cartesian
-- closed will follow immediately from these constructions.
-- | the category of categories have products.
module CatProduct { ' : Level} ( 𝔻 : Category ') where
private
2018-03-08 10:54:13 +00:00
module = Category
module 𝔻 = Category 𝔻
Obj = Object × Object 𝔻
Arr : Obj Obj Set '
Arr (c , d) (c' , d') = [ c , c' ] × 𝔻 [ d , d' ]
𝟙' : {o : Obj} Arr o o
𝟙' = 𝟙 , 𝟙 𝔻
_∘_ :
{a b c : Obj}
Arr b c
Arr a b
Arr a c
_∘_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) [ bc∈C ab∈C ] , 𝔻 [ bc∈D ab∈D ]}
rawProduct : RawCategory '
RawCategory.Object rawProduct = Obj
RawCategory.Arrow rawProduct = Arr
RawCategory.𝟙 rawProduct = 𝟙'
RawCategory._∘_ rawProduct = _∘_
open RawCategory rawProduct
arrowsAreSets : ArrowsAreSets
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
2018-03-08 10:54:13 +00:00
object : Category '
Category.raw object = rawProduct
2018-03-08 10:54:13 +00:00
proj₁ : Functor object
proj₁ = record
2018-03-08 10:54:13 +00:00
{ raw = record
{ omap = fst ; fmap = fst }
; isFunctor = record
{ isIdentity = refl ; isDistributive = refl }
}
2018-03-08 10:54:13 +00:00
proj₂ : Functor object 𝔻
proj₂ = record
2018-03-08 10:54:13 +00:00
{ raw = record
{ omap = snd ; fmap = 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
2018-03-08 10:54:13 +00:00
x : Functor X object
x = record
{ raw = record
{ omap = λ x x₁.omap x , x₂.omap x
; fmap = λ x x₁.fmap x , x₂.fmap x
}
; isFunctor = record
2018-03-08 10:54:13 +00:00
{ 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 10:54:13 +00:00
RawProduct.object rawProduct = P.object
2018-03-08 09:45:15 +00:00
RawProduct.proj₁ rawProduct = P.proj₁
RawProduct.proj₂ rawProduct = P.proj₂
2018-03-08 09:20:29 +00:00
2018-03-08 09:28:05 +00:00
isProduct : IsProduct Cat _ _ rawProduct
2018-03-08 09:20:29 +00:00
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-03-08 10:54:13 +00:00
-- | The category of categories have expoentntials - and because it has products
-- it is therefory also cartesian closed.
module CatExponential { : Level} ( 𝔻 : Category ) where
2018-03-08 10:20:51 +00:00
private
open Data.Product
open import Cat.Categories.Fun
module = Category
module 𝔻 = Category 𝔻
2018-03-08 10:54:13 +00:00
Category = Category
open Fun 𝔻 renaming (identity to idN)
omap : Functor 𝔻 × Object Object 𝔻
2018-03-08 10:54:13 +00:00
omap (F , A) = Functor.omap F A
2018-01-25 11:01:37 +00:00
2018-03-08 10:54:13 +00:00
-- The exponential object
object : Category
object = 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
2018-03-08 10:20:51 +00:00
module F = Functor F
module G = Functor G
fmap : (pobj : NaturalTransformation F G × [ A , B ])
2018-03-08 10:20:51 +00:00
𝔻 [ F.omap A , G.omap B ]
fmap ((θ , θNat) , f) = result
where
2018-03-08 10:20:51 +00:00
θA : 𝔻 [ F.omap A , G.omap A ]
θA = θ A
2018-03-08 10:20:51 +00:00
θB : 𝔻 [ F.omap B , G.omap B ]
θB = θ B
2018-03-08 10:20:51 +00:00
F→f : 𝔻 [ F.omap A , F.omap B ]
F→f = F.fmap f
G→f : 𝔻 [ G.omap A , G.omap B ]
G→f = G.fmap f
l : 𝔻 [ F.omap A , G.omap B ]
l = 𝔻 [ θB F.fmap f ]
r : 𝔻 [ F.omap A , G.omap B ]
r = 𝔻 [ G.fmap f θA ]
result : 𝔻 [ F.omap A , G.omap B ]
result = l
2018-03-08 10:54:13 +00:00
open CatProduct renaming (object to _⊗_) using ()
module _ {c : Functor 𝔻 × Object } where
private
F : Functor 𝔻
F = proj₁ c
C : Object
C = proj₂ c
ident : fmap {c} {c} (NT.identity F , 𝟙 {A = proj₂ c}) 𝟙 𝔻
ident = begin
2018-03-08 10:54:13 +00:00
fmap {c} {c} (𝟙 (object ) {c}) ≡⟨⟩
fmap {c} {c} (idN F , 𝟙 ) ≡⟨⟩
2018-03-08 10:20:51 +00:00
𝔻 [ identityTrans F C F.fmap (𝟙 )] ≡⟨⟩
𝔻 [ 𝟙 𝔻 F.fmap (𝟙 )] ≡⟨ proj₂ 𝔻.isIdentity
F.fmap (𝟙 ) ≡⟨ F.isIdentity
𝟙 𝔻
where
module F = Functor F
module _ {F×A G×B H×C : Functor 𝔻 × Object } where
2018-03-08 10:20:51 +00:00
private
F = F×A .proj₁
A = F×A .proj₂
G = G×B .proj₁
B = G×B .proj₂
H = H×C .proj₁
C = H×C .proj₂
module F = Functor F
module G = Functor G
module H = Functor H
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 :
2018-03-08 10:20:51 +00:00
𝔻 [ 𝔻 [ η C θ C ] F.fmap ( [ g f ] ) ]
𝔻 [ 𝔻 [ η C G.fmap g ] 𝔻 [ θ B F.fmap f ] ]
isDistributive = begin
2018-03-08 10:20:51 +00:00
𝔻 [ (ηθ C) F.fmap ( [ g f ]) ]
≡⟨ ηθNat ( [ g f ])
2018-03-08 10:20:51 +00:00
𝔻 [ H.fmap ( [ g f ]) (ηθ A) ]
≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.isDistributive)
2018-03-08 10:20:51 +00:00
𝔻 [ 𝔻 [ H.fmap g H.fmap f ] (ηθ A) ]
≡⟨ sym 𝔻.isAssociative
𝔻 [ H.fmap g 𝔻 [ H.fmap f ηθ A ] ]
≡⟨ cong (λ φ 𝔻 [ H.fmap g φ ]) 𝔻.isAssociative
𝔻 [ H.fmap g 𝔻 [ 𝔻 [ H.fmap f η A ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ H.fmap g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f)))
𝔻 [ H.fmap g 𝔻 [ 𝔻 [ η B G.fmap f ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ H.fmap g φ ]) (sym 𝔻.isAssociative)
𝔻 [ H.fmap g 𝔻 [ η B 𝔻 [ G.fmap f θ A ] ] ]
≡⟨ 𝔻.isAssociative
𝔻 [ 𝔻 [ H.fmap g η B ] 𝔻 [ G.fmap f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ G.fmap f θ A ] ]) (sym (ηNat g))
𝔻 [ 𝔻 [ η C G.fmap g ] 𝔻 [ G.fmap f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C G.fmap g ] φ ]) (sym (θNat f))
𝔻 [ 𝔻 [ η C G.fmap g ] 𝔻 [ θ B F.fmap f ] ]
2018-03-08 10:54:13 +00:00
eval : Functor (CatProduct.object object ) 𝔻
eval = record
{ raw = record
{ omap = omap
; fmap = λ {dom} {cod} fmap {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
2018-03-08 10:54:13 +00:00
module _ (𝔸 : Category ) (F : Functor (𝔸 ) 𝔻) where
postulate
parallelProduct
2018-03-08 10:54:13 +00:00
: Functor 𝔸 object Functor
Functor (𝔸 ) (object )
transpose : Functor 𝔸 object
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
2018-03-08 10:54:13 +00:00
-- We don't care about filling out the holes below since they are anyways hidden
-- behind an unprovable statement.
module _ ( : Level) (unprovable : IsCategory (RawCat )) where
private
Cat : Category (lsuc ( )) ( )
Cat = Cat unprovable
2018-03-08 10:54:13 +00:00
module _ ( 𝔻 : Category ) where
module CatExp = CatExponential 𝔻
_⊗_ = CatProduct.object
-- Filling the hole causes Agda to loop indefinitely.
eval : Functor (CatExp.object ) 𝔻
eval = {!CatExp.eval!}
isExponential : IsExponential Cat 𝔻 CatExp.object eval
isExponential = {!CatExp.isExponential!}
exponent : Exponential Cat 𝔻
exponent = record
{ obj = CatExp.object
; eval = eval
; isExponential = isExponential
}
2018-02-23 09:44:23 +00:00
hasExponentials : HasExponentials Cat
hasExponentials = record { exponent = exponent }