2018-02-02 14:33:54 +00:00
|
|
|
|
-- There is no category of categories in our interpretation
|
2018-01-17 22:00:27 +00:00
|
|
|
|
{-# OPTIONS --cubical --allow-unsolved-metas #-}
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
2018-01-17 22:00:27 +00:00
|
|
|
|
module Cat.Categories.Cat where
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
2018-04-05 08:41:56 +00:00
|
|
|
|
open import Cat.Prelude renaming (fst to fst ; snd to snd)
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
2018-01-17 22:00:27 +00:00
|
|
|
|
open import Cat.Category
|
2018-02-05 15:35:33 +00:00
|
|
|
|
open import Cat.Category.Functor
|
|
|
|
|
open import Cat.Category.Product
|
2018-03-05 12:52:41 +00:00
|
|
|
|
open import Cat.Category.Exponential hiding (_×_ ; product)
|
2018-03-23 14:20:26 +00:00
|
|
|
|
import Cat.Category.NaturalTransformation
|
2018-03-21 13:39:56 +00:00
|
|
|
|
open import Cat.Categories.Fun
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
|
|
|
|
-- The category of categories
|
2018-01-24 15:38:28 +00:00
|
|
|
|
module _ (ℓ ℓ' : Level) where
|
2018-02-23 09:34:37 +00:00
|
|
|
|
RawCat : RawCategory (lsuc (ℓ ⊔ ℓ')) (ℓ ⊔ ℓ')
|
2018-04-03 09:36:09 +00:00
|
|
|
|
RawCategory.Object RawCat = Category ℓ ℓ'
|
|
|
|
|
RawCategory.Arrow RawCat = Functor
|
|
|
|
|
RawCategory.identity RawCat = Functors.identity
|
|
|
|
|
RawCategory._∘_ RawCat = F[_∘_]
|
2018-03-21 12:25:24 +00:00
|
|
|
|
|
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.
|
2018-02-23 09:34:37 +00:00
|
|
|
|
Cat : (unprovable : IsCategory RawCat) → Category (lsuc (ℓ ⊔ ℓ')) (ℓ ⊔ ℓ')
|
2018-03-21 12:25:24 +00:00
|
|
|
|
Category.raw (Cat _) = RawCat
|
2018-02-23 09:34:37 +00:00
|
|
|
|
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.
|
2018-03-05 10:13:37 +00:00
|
|
|
|
module CatProduct {ℓ ℓ' : Level} (ℂ 𝔻 : Category ℓ ℓ') where
|
2018-03-05 10:01:36 +00:00
|
|
|
|
private
|
2018-03-08 10:54:13 +00:00
|
|
|
|
module ℂ = Category ℂ
|
|
|
|
|
module 𝔻 = Category 𝔻
|
|
|
|
|
|
2018-03-21 12:25:24 +00:00
|
|
|
|
module _ where
|
|
|
|
|
private
|
|
|
|
|
Obj = ℂ.Object × 𝔻.Object
|
|
|
|
|
Arr : Obj → Obj → Set ℓ'
|
|
|
|
|
Arr (c , d) (c' , d') = ℂ [ c , c' ] × 𝔻 [ d , d' ]
|
2018-04-03 09:36:09 +00:00
|
|
|
|
identity : {o : Obj} → Arr o o
|
|
|
|
|
identity = ℂ.identity , 𝔻.identity
|
2018-03-21 12:25:24 +00:00
|
|
|
|
_∘_ :
|
|
|
|
|
{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 ℓ ℓ'
|
2018-04-03 09:36:09 +00:00
|
|
|
|
RawCategory.Object rawProduct = Obj
|
|
|
|
|
RawCategory.Arrow rawProduct = Arr
|
|
|
|
|
RawCategory.identity rawProduct = identity
|
|
|
|
|
RawCategory._∘_ rawProduct = _∘_
|
2018-03-21 12:25:24 +00:00
|
|
|
|
|
|
|
|
|
open RawCategory rawProduct
|
2018-03-05 10:01:36 +00:00
|
|
|
|
|
2018-03-08 10:29:16 +00:00
|
|
|
|
arrowsAreSets : ArrowsAreSets
|
2018-03-05 10:01:36 +00:00
|
|
|
|
arrowsAreSets = setSig {sA = ℂ.arrowsAreSets} {sB = λ x → 𝔻.arrowsAreSets}
|
2018-04-03 09:36:09 +00:00
|
|
|
|
isIdentity : IsIdentity identity
|
2018-03-05 10:01:36 +00:00
|
|
|
|
isIdentity
|
|
|
|
|
= Σ≡ (fst ℂ.isIdentity) (fst 𝔻.isIdentity)
|
|
|
|
|
, Σ≡ (snd ℂ.isIdentity) (snd 𝔻.isIdentity)
|
2018-03-20 14:19:28 +00:00
|
|
|
|
|
|
|
|
|
postulate univalent : Univalence.Univalent isIdentity
|
2018-03-05 10:01:36 +00:00
|
|
|
|
instance
|
2018-03-08 10:29:16 +00:00
|
|
|
|
isCategory : IsCategory rawProduct
|
|
|
|
|
IsCategory.isAssociative isCategory = Σ≡ ℂ.isAssociative 𝔻.isAssociative
|
|
|
|
|
IsCategory.isIdentity isCategory = isIdentity
|
|
|
|
|
IsCategory.arrowsAreSets isCategory = arrowsAreSets
|
|
|
|
|
IsCategory.univalent isCategory = univalent
|
2018-03-05 10:01:36 +00:00
|
|
|
|
|
2018-03-08 10:54:13 +00:00
|
|
|
|
object : Category ℓ ℓ'
|
|
|
|
|
Category.raw object = rawProduct
|
2018-03-05 10:01:36 +00:00
|
|
|
|
|
2018-04-05 08:41:56 +00:00
|
|
|
|
fstF : Functor object ℂ
|
|
|
|
|
fstF = record
|
2018-03-08 10:54:13 +00:00
|
|
|
|
{ raw = record
|
|
|
|
|
{ omap = fst ; fmap = fst }
|
|
|
|
|
; isFunctor = record
|
|
|
|
|
{ isIdentity = refl ; isDistributive = refl }
|
2018-03-05 10:01:36 +00:00
|
|
|
|
}
|
|
|
|
|
|
2018-04-05 08:41:56 +00:00
|
|
|
|
sndF : Functor object 𝔻
|
|
|
|
|
sndF = record
|
2018-03-08 10:54:13 +00:00
|
|
|
|
{ raw = record
|
|
|
|
|
{ omap = snd ; fmap = snd }
|
|
|
|
|
; isFunctor = record
|
|
|
|
|
{ isIdentity = refl ; isDistributive = refl }
|
2018-03-05 10:01:36 +00:00
|
|
|
|
}
|
|
|
|
|
|
|
|
|
|
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
|
2018-03-05 10:01:36 +00:00
|
|
|
|
x = record
|
|
|
|
|
{ raw = record
|
2018-03-08 10:03:56 +00:00
|
|
|
|
{ omap = λ x → x₁.omap x , x₂.omap x
|
|
|
|
|
; fmap = λ x → x₁.fmap x , x₂.fmap x
|
2018-02-23 11:29:10 +00:00
|
|
|
|
}
|
2018-03-05 10:01:36 +00:00
|
|
|
|
; isFunctor = record
|
2018-03-08 10:54:13 +00:00
|
|
|
|
{ isIdentity = Σ≡ x₁.isIdentity x₂.isIdentity
|
2018-03-05 10:01:36 +00:00
|
|
|
|
; isDistributive = Σ≡ x₁.isDistributive x₂.isDistributive
|
|
|
|
|
}
|
|
|
|
|
}
|
|
|
|
|
where
|
|
|
|
|
open module x₁ = Functor x₁
|
|
|
|
|
open module x₂ = Functor x₂
|
2018-02-23 11:29:10 +00:00
|
|
|
|
|
2018-04-05 08:41:56 +00:00
|
|
|
|
isUniqL : F[ fstF ∘ x ] ≡ x₁
|
2018-03-05 16:10:41 +00:00
|
|
|
|
isUniqL = Functor≡ refl
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-04-05 08:41:56 +00:00
|
|
|
|
isUniqR : F[ sndF ∘ x ] ≡ x₂
|
2018-03-05 16:10:41 +00:00
|
|
|
|
isUniqR = Functor≡ refl
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-04-05 08:41:56 +00:00
|
|
|
|
isUniq : F[ fstF ∘ x ] ≡ x₁ × F[ sndF ∘ x ] ≡ x₂
|
2018-03-05 10:01:36 +00:00
|
|
|
|
isUniq = isUniqL , isUniqR
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-04-05 08:41:56 +00:00
|
|
|
|
isProduct : ∃![ x ] (F[ fstF ∘ x ] ≡ x₁ × F[ sndF ∘ x ] ≡ x₂)
|
2018-03-27 12:18:13 +00:00
|
|
|
|
isProduct = x , isUniq , uq
|
|
|
|
|
where
|
2018-04-05 08:41:56 +00:00
|
|
|
|
module _ {y : Functor X object} (eq : F[ fstF ∘ y ] ≡ x₁ × F[ sndF ∘ y ] ≡ x₂) where
|
2018-03-27 12:18:13 +00:00
|
|
|
|
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
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-03-05 10:01:36 +00:00
|
|
|
|
module _ {ℓ ℓ' : Level} (unprovable : IsCategory (RawCat ℓ ℓ')) where
|
|
|
|
|
private
|
|
|
|
|
Catℓ = Cat ℓ ℓ' unprovable
|
2018-03-05 10:13:37 +00:00
|
|
|
|
|
2018-03-05 10:01:36 +00:00
|
|
|
|
module _ (ℂ 𝔻 : Category ℓ ℓ') where
|
|
|
|
|
private
|
2018-03-05 10:13:37 +00:00
|
|
|
|
module P = CatProduct ℂ 𝔻
|
2018-03-05 09:35:33 +00:00
|
|
|
|
|
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-04-05 08:41:56 +00:00
|
|
|
|
RawProduct.fst rawProduct = P.fstF
|
|
|
|
|
RawProduct.snd rawProduct = P.sndF
|
2018-03-08 09:20:29 +00:00
|
|
|
|
|
2018-03-08 09:28:05 +00:00
|
|
|
|
isProduct : IsProduct Catℓ _ _ rawProduct
|
2018-03-14 09:23:23 +00:00
|
|
|
|
IsProduct.ump 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-21 13:31:37 +00:00
|
|
|
|
|
2018-01-24 15:38:28 +00:00
|
|
|
|
instance
|
2018-03-05 10:01:36 +00:00
|
|
|
|
hasProducts : HasProducts Catℓ
|
|
|
|
|
hasProducts = record { product = product }
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
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.
|
2018-03-05 12:52:41 +00:00
|
|
|
|
module CatExponential {ℓ : Level} (ℂ 𝔻 : Category ℓ ℓ) where
|
2018-03-23 14:20:26 +00:00
|
|
|
|
open Cat.Category.NaturalTransformation ℂ 𝔻
|
|
|
|
|
renaming (identity to identityNT)
|
|
|
|
|
using ()
|
2018-03-08 10:20:51 +00:00
|
|
|
|
private
|
|
|
|
|
module ℂ = Category ℂ
|
|
|
|
|
module 𝔻 = Category 𝔻
|
2018-03-08 10:54:13 +00:00
|
|
|
|
Categoryℓ = Category ℓ ℓ
|
|
|
|
|
open Fun ℂ 𝔻 renaming (identity to idN)
|
2018-03-05 12:52:41 +00:00
|
|
|
|
|
2018-03-21 12:25:24 +00:00
|
|
|
|
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
|
2018-03-05 12:52:41 +00:00
|
|
|
|
|
2018-03-21 12:25:24 +00:00
|
|
|
|
module _ {dom cod : Functor ℂ 𝔻 × ℂ.Object} where
|
2018-04-05 08:41:56 +00:00
|
|
|
|
open Σ dom renaming (fst to F ; snd to A)
|
|
|
|
|
open Σ cod renaming (fst to G ; snd to B)
|
2018-03-05 12:52:41 +00:00
|
|
|
|
private
|
2018-03-08 10:20:51 +00:00
|
|
|
|
module F = Functor F
|
|
|
|
|
module G = Functor G
|
|
|
|
|
|
2018-03-08 10:29:16 +00:00
|
|
|
|
fmap : (pobj : NaturalTransformation F G × ℂ [ A , B ])
|
2018-03-08 10:20:51 +00:00
|
|
|
|
→ 𝔻 [ F.omap A , G.omap B ]
|
2018-03-21 12:25:24 +00:00
|
|
|
|
fmap ((θ , θNat) , f) = 𝔻 [ θ B ∘ F.fmap f ]
|
|
|
|
|
-- Alternatively:
|
|
|
|
|
--
|
|
|
|
|
-- fmap ((θ , θNat) , f) = 𝔻 [ G.fmap f ∘ θ A ]
|
|
|
|
|
--
|
|
|
|
|
-- Since they are equal by naturality of θ.
|
2018-03-05 12:52:41 +00:00
|
|
|
|
|
2018-03-08 10:54:13 +00:00
|
|
|
|
open CatProduct renaming (object to _⊗_) using ()
|
2018-03-05 12:52:41 +00:00
|
|
|
|
|
2018-03-21 12:25:24 +00:00
|
|
|
|
module _ {c : Functor ℂ 𝔻 × ℂ.Object} where
|
2018-04-05 08:41:56 +00:00
|
|
|
|
open Σ c renaming (fst to F ; snd to C)
|
2018-03-05 12:52:41 +00:00
|
|
|
|
|
2018-04-03 09:36:09 +00:00
|
|
|
|
ident : fmap {c} {c} (identityNT F , ℂ.identity {A = snd c}) ≡ 𝔻.identity
|
2018-03-08 10:29:16 +00:00
|
|
|
|
ident = begin
|
2018-04-03 09:36:09 +00:00
|
|
|
|
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 ∎
|
2018-03-05 12:52:41 +00:00
|
|
|
|
where
|
2018-03-08 10:29:16 +00:00
|
|
|
|
module F = Functor F
|
2018-03-05 12:52:41 +00:00
|
|
|
|
|
2018-03-21 12:25:24 +00:00
|
|
|
|
module _ {F×A G×B H×C : Functor ℂ 𝔻 × ℂ.Object} where
|
2018-04-05 08:41:56 +00:00
|
|
|
|
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)
|
2018-03-08 10:20:51 +00:00
|
|
|
|
private
|
|
|
|
|
module F = Functor F
|
|
|
|
|
module G = Functor G
|
|
|
|
|
module H = Functor H
|
|
|
|
|
|
2018-03-05 12:52:41 +00:00
|
|
|
|
module _
|
|
|
|
|
{θ×f : NaturalTransformation F G × ℂ [ A , B ]}
|
|
|
|
|
{η×g : NaturalTransformation G H × ℂ [ B , C ]} where
|
2018-04-05 08:41:56 +00:00
|
|
|
|
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)
|
2018-01-24 15:38:28 +00:00
|
|
|
|
private
|
2018-03-05 12:52:41 +00:00
|
|
|
|
ηθNT : NaturalTransformation F H
|
2018-03-21 12:25:24 +00:00
|
|
|
|
ηθNT = NT[_∘_] {F} {G} {H} ηNT θNT
|
2018-04-05 08:41:56 +00:00
|
|
|
|
open Σ ηθNT renaming (fst to ηθ ; snd to ηθNat)
|
2018-03-05 12:52:41 +00:00
|
|
|
|
|
2018-03-08 10:29:16 +00:00
|
|
|
|
isDistributive :
|
2018-03-08 10:20:51 +00:00
|
|
|
|
𝔻 [ 𝔻 [ η C ∘ θ C ] ∘ F.fmap ( ℂ [ g ∘ f ] ) ]
|
|
|
|
|
≡ 𝔻 [ 𝔻 [ η C ∘ G.fmap g ] ∘ 𝔻 [ θ B ∘ F.fmap f ] ]
|
2018-03-08 10:29:16 +00:00
|
|
|
|
isDistributive = begin
|
2018-03-08 10:20:51 +00:00
|
|
|
|
𝔻 [ (ηθ C) ∘ F.fmap (ℂ [ g ∘ f ]) ]
|
2018-03-05 12:52:41 +00:00
|
|
|
|
≡⟨ ηθNat (ℂ [ g ∘ f ]) ⟩
|
2018-03-08 10:20:51 +00:00
|
|
|
|
𝔻 [ H.fmap (ℂ [ g ∘ f ]) ∘ (ηθ A) ]
|
2018-03-05 12:52:41 +00:00
|
|
|
|
≡⟨ 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-05 12:52:41 +00:00
|
|
|
|
|
2018-03-08 10:54:13 +00:00
|
|
|
|
eval : Functor (CatProduct.object object ℂ) 𝔻
|
2018-03-05 12:52:41 +00:00
|
|
|
|
eval = record
|
|
|
|
|
{ raw = record
|
2018-03-08 10:29:16 +00:00
|
|
|
|
{ omap = omap
|
|
|
|
|
; fmap = λ {dom} {cod} → fmap {dom} {cod}
|
2018-03-05 12:52:41 +00:00
|
|
|
|
}
|
|
|
|
|
; isFunctor = record
|
2018-03-08 10:29:16 +00:00
|
|
|
|
{ isIdentity = λ {o} → ident {o}
|
|
|
|
|
; isDistributive = λ {f u n k y} → isDistributive {f} {u} {n} {k} {y}
|
2018-03-05 12:52:41 +00:00
|
|
|
|
}
|
|
|
|
|
}
|
2018-02-23 09:44:23 +00:00
|
|
|
|
|
2018-03-08 10:54:13 +00:00
|
|
|
|
module _ (𝔸 : Category ℓ ℓ) (F : Functor (𝔸 ⊗ ℂ) 𝔻) where
|
2018-03-05 12:52:41 +00:00
|
|
|
|
postulate
|
|
|
|
|
parallelProduct
|
2018-03-08 10:54:13 +00:00
|
|
|
|
: Functor 𝔸 object → Functor ℂ ℂ
|
|
|
|
|
→ Functor (𝔸 ⊗ ℂ) (object ⊗ ℂ)
|
|
|
|
|
transpose : Functor 𝔸 object
|
2018-03-23 12:55:03 +00:00
|
|
|
|
eq : F[ eval ∘ (parallelProduct transpose (Functors.identity {ℂ = ℂ})) ] ≡ F
|
2018-03-05 12:52:41 +00:00
|
|
|
|
-- eq : F[ :eval: ∘ {!!} ] ≡ F
|
2018-04-03 09:36:09 +00:00
|
|
|
|
-- eq : Catℓ [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (identity Catℓ {o = ℂ})) ] ≡ F
|
2018-03-05 12:52:41 +00:00
|
|
|
|
-- eq' : (Catℓ [ :eval: ∘
|
|
|
|
|
-- (record { product = product } HasProducts.|×| transpose)
|
2018-04-03 09:36:09 +00:00
|
|
|
|
-- (identity Catℓ)
|
2018-03-05 12:52:41 +00:00
|
|
|
|
-- ])
|
|
|
|
|
-- ≡ F
|
|
|
|
|
|
|
|
|
|
-- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758`
|
|
|
|
|
-- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Catℓ [
|
2018-04-03 09:36:09 +00:00
|
|
|
|
-- :eval: ∘ (parallelProduct F~ (identity Catℓ {o = ℂ}))] ≡ F) catTranspose =
|
2018-03-05 12:52:41 +00:00
|
|
|
|
-- 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.
|
2018-03-05 12:52:41 +00:00
|
|
|
|
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ℓ
|
2018-03-05 12:52:41 +00:00
|
|
|
|
hasExponentials = record { exponent = exponent }
|