cat/src/Cat/Categories/Cat.agda

324 lines
12 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

-- There is no category of categories in our interpretation
{-# OPTIONS --cubical --allow-unsolved-metas #-}
module Cat.Categories.Cat where
open import Cat.Prelude renaming (fst to fst ; snd to snd)
open import Cat.Category
open import Cat.Category.Functor
open import Cat.Category.Product
open import Cat.Category.Exponential hiding (_×_ ; product)
import Cat.Category.NaturalTransformation
open import Cat.Categories.Fun
-- The category of categories
module _ ( ' : Level) where
RawCat : RawCategory (lsuc ( ')) ( ')
RawCategory.Object RawCat = Category '
RawCategory.Arrow RawCat = Functor
RawCategory.identity RawCat = Functors.identity
RawCategory._<<<_ RawCat = F[_∘_]
-- 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
-- | 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
module = Category
module 𝔻 = Category 𝔻
module _ where
private
Obj = .Object × 𝔻.Object
Arr : Obj Obj Set '
Arr (c , d) (c' , d') = [ c , c' ] × 𝔻 [ d , d' ]
identity : {o : Obj} Arr o o
identity = .identity , 𝔻.identity
_<<<_ :
{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.identity rawProduct = identity
RawCategory._<<<_ rawProduct = _<<<_
open RawCategory rawProduct
arrowsAreSets : ArrowsAreSets
arrowsAreSets = setSig {sA = .arrowsAreSets} {sB = λ x 𝔻.arrowsAreSets}
isIdentity : IsIdentity identity
isIdentity
= Σ≡ (fst .isIdentity) (fst 𝔻.isIdentity)
, Σ≡ (snd .isIdentity) (snd 𝔻.isIdentity)
isPreCategory : IsPreCategory rawProduct
IsPreCategory.isAssociative isPreCategory = Σ≡ .isAssociative 𝔻.isAssociative
IsPreCategory.isIdentity isPreCategory = isIdentity
IsPreCategory.arrowsAreSets isPreCategory = arrowsAreSets
postulate univalent : Univalence.Univalent isIdentity
isCategory : IsCategory rawProduct
IsCategory.isPreCategory isCategory = isPreCategory
IsCategory.univalent isCategory = univalent
object : Category '
Category.raw object = rawProduct
Category.isCategory object = isCategory
fstF : Functor object
fstF = record
{ raw = record
{ omap = fst ; fmap = fst }
; isFunctor = record
{ isIdentity = refl ; isDistributive = refl }
}
sndF : Functor object 𝔻
sndF = record
{ raw = record
{ omap = snd ; fmap = snd }
; isFunctor = record
{ isIdentity = refl ; isDistributive = refl }
}
module _ {X : Category '} (x₁ : Functor X ) (x₂ : Functor X 𝔻) where
private
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
{ isIdentity = Σ≡ x₁.isIdentity x₂.isIdentity
; isDistributive = Σ≡ x₁.isDistributive x₂.isDistributive
}
}
where
open module x = Functor x₁
open module x = Functor x₂
isUniqL : F[ fstF x ] x₁
isUniqL = Functor≡ refl
isUniqR : F[ sndF x ] x₂
isUniqR = Functor≡ refl
isUniq : F[ fstF x ] x₁ × F[ sndF x ] x₂
isUniq = isUniqL , isUniqR
isProduct : ∃![ x ] (F[ fstF x ] x₁ × F[ sndF x ] x₂)
isProduct = x , isUniq , uq
where
module _ {y : Functor X object} (eq : F[ fstF y ] x₁ × F[ sndF y ] x₂) where
omapEq : Functor.omap x Functor.omap y
omapEq = {!!}
-- fmapEq : (λ i → {!{A B : ?} → Arrow A B → 𝔻 [ ? A , ? B ]!}) [ Functor.fmap x ≡ Functor.fmap y ]
-- fmapEq = {!!}
rawEq : Functor.raw x Functor.raw y
rawEq = {!!}
uq : x y
uq = Functor≡ rawEq
module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
private
Cat = Cat ' unprovable
module _ ( 𝔻 : Category ') where
private
module P = CatProduct 𝔻
rawProduct : RawProduct Cat 𝔻
RawProduct.object rawProduct = P.object
RawProduct.fst rawProduct = P.fstF
RawProduct.snd rawProduct = P.sndF
isProduct : IsProduct Cat _ _ rawProduct
IsProduct.ump isProduct = P.isProduct
product : Product Cat 𝔻
Product.raw product = rawProduct
Product.isProduct product = isProduct
instance
hasProducts : HasProducts Cat
hasProducts = record { product = product }
-- | The category of categories have expoentntials - and because it has products
-- it is therefory also cartesian closed.
module CatExponential { : Level} ( 𝔻 : Category ) where
open Cat.Category.NaturalTransformation 𝔻
renaming (identity to identityNT)
using ()
private
module = Category
module 𝔻 = Category 𝔻
Category = Category
open Fun 𝔻 renaming (identity to idN)
omap : Functor 𝔻 × .Object 𝔻.Object
omap (F , A) = Functor.omap F A
-- The exponential object
object : Category
object = Fun
module _ {dom cod : Functor 𝔻 × .Object} where
open Σ dom renaming (fst to F ; snd to A)
open Σ cod renaming (fst to G ; snd to B)
private
module F = Functor F
module G = Functor G
fmap : (pobj : NaturalTransformation F G × [ A , B ])
𝔻 [ F.omap A , G.omap B ]
fmap ((θ , θNat) , f) = 𝔻 [ θ B F.fmap f ]
-- Alternatively:
--
-- fmap ((θ , θNat) , f) = 𝔻 [ G.fmap f ∘ θ A ]
--
-- Since they are equal by naturality of θ.
open CatProduct renaming (object to _⊗_) using ()
module _ {c : Functor 𝔻 × .Object} where
open Σ c renaming (fst to F ; snd to C)
ident : fmap {c} {c} (identityNT F , .identity {A = snd c}) 𝔻.identity
ident = begin
fmap {c} {c} (Category.identity (object ) {c}) ≡⟨⟩
fmap {c} {c} (idN F , .identity) ≡⟨⟩
𝔻 [ identityTrans F C F.fmap .identity ] ≡⟨⟩
𝔻 [ 𝔻.identity F.fmap .identity ] ≡⟨ 𝔻.leftIdentity
F.fmap .identity ≡⟨ F.isIdentity
𝔻.identity
where
module F = Functor F
module _ {F×A G×B H×C : Functor 𝔻 × .Object} where
open Σ F×A renaming (fst to F ; snd to A)
open Σ G×B renaming (fst to G ; snd to B)
open Σ H×C renaming (fst to H ; snd to C)
private
module F = Functor F
module G = Functor G
module H = Functor H
module _
{θ×f : NaturalTransformation F G × [ A , B ]}
{η×g : NaturalTransformation G H × [ B , C ]} where
open Σ θ×f renaming (fst to θNT ; snd to f)
open Σ θNT renaming (fst to θ ; snd to θNat)
open Σ η×g renaming (fst to ηNT ; snd to g)
open Σ ηNT renaming (fst to η ; snd to ηNat)
private
ηθNT : NaturalTransformation F H
ηθNT = NT[_∘_] {F} {G} {H} ηNT θNT
open Σ ηθNT renaming (fst to ηθ ; snd to ηθNat)
isDistributive :
𝔻 [ 𝔻 [ η C θ C ] F.fmap ( [ g f ] ) ]
𝔻 [ 𝔻 [ η C G.fmap g ] 𝔻 [ θ B F.fmap f ] ]
isDistributive = begin
𝔻 [ (ηθ C) F.fmap ( [ g f ]) ]
≡⟨ ηθNat ( [ g f ])
𝔻 [ H.fmap ( [ g f ]) (ηθ A) ]
≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.isDistributive)
𝔻 [ 𝔻 [ 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 ] ]
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}
}
}
module _ (𝔸 : Category ) (F : Functor (𝔸 ) 𝔻) where
postulate
parallelProduct
: Functor 𝔸 object Functor
Functor (𝔸 ) (object )
transpose : Functor 𝔸 object
eq : F[ eval (parallelProduct transpose (Functors.identity { = })) ] F
-- eq : F[ :eval: ∘ {!!} ] ≡ F
-- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (identity Cat {o = })) ] ≡ F
-- eq' : (Cat [ :eval: ∘
-- (record { product = product } HasProducts.|×| transpose)
-- (identity Cat)
-- ])
-- ≡ F
-- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758`
-- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Cat [
-- :eval: (parallelProduct F~ (identity Cat {o = }))] F) catTranspose =
-- transpose , eq
-- 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
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
}
hasExponentials : HasExponentials Cat
hasExponentials = record { exponent = exponent }