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
|
|
|
|
|
|
|
|
|
open import Agda.Primitive
|
|
|
|
|
open import Cubical
|
|
|
|
|
open import Function
|
|
|
|
|
open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
|
|
|
|
|
|
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
|
|
|
|
|
open import Cat.Category.Exponential
|
2018-02-23 16:33:09 +00:00
|
|
|
|
open import Cat.Category.NaturalTransformation
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
2018-01-31 13:39:54 +00:00
|
|
|
|
open import Cat.Equality
|
|
|
|
|
open Equality.Data.Product
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
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
|
|
|
|
|
2018-01-08 21:29:29 +00:00
|
|
|
|
-- The category of categories
|
2018-01-24 15:38:28 +00:00
|
|
|
|
module _ (ℓ ℓ' : Level) where
|
2018-01-08 21:29:29 +00:00
|
|
|
|
private
|
2018-01-31 13:39:54 +00:00
|
|
|
|
module _ {𝔸 𝔹 ℂ 𝔻 : Category ℓ ℓ'} {F : Functor 𝔸 𝔹} {G : Functor 𝔹 ℂ} {H : Functor ℂ 𝔻} where
|
|
|
|
|
assc : H ∘f (G ∘f F) ≡ (H ∘f G) ∘f F
|
2018-02-23 11:15:16 +00:00
|
|
|
|
assc = Functor≡ refl refl
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
2018-01-25 11:44:47 +00:00
|
|
|
|
module _ {ℂ 𝔻 : Category ℓ ℓ'} {F : Functor ℂ 𝔻} where
|
2018-02-23 11:15:16 +00:00
|
|
|
|
ident-r : F ∘f identity ≡ F
|
|
|
|
|
ident-r = Functor≡ refl refl
|
|
|
|
|
|
|
|
|
|
ident-l : identity ∘f F ≡ F
|
|
|
|
|
ident-l = Functor≡ refl refl
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
2018-02-23 09:34:37 +00:00
|
|
|
|
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-02-23 09:34:37 +00:00
|
|
|
|
-- TODO: Rename `ident'` to `ident` after changing how names are exposed in Functor.
|
2018-02-23 11:15:16 +00:00
|
|
|
|
ident' : IsIdentity identity
|
2018-02-23 09:34:37 +00:00
|
|
|
|
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 _ {ℓ ℓ' : Level} (unprovable : IsCategory (RawCat ℓ ℓ')) where
|
2018-01-25 11:44:47 +00:00
|
|
|
|
module _ (ℂ 𝔻 : Category ℓ ℓ') where
|
2018-01-24 15:38:28 +00:00
|
|
|
|
private
|
2018-02-23 09:34:37 +00:00
|
|
|
|
Catt = Cat ℓ ℓ' unprovable
|
2018-02-05 15:35:33 +00:00
|
|
|
|
:Object: = Object ℂ × Object 𝔻
|
2018-01-24 15:38:28 +00:00
|
|
|
|
:Arrow: : :Object: → :Object: → Set ℓ'
|
2018-02-22 14:31:54 +00:00
|
|
|
|
:Arrow: (c , d) (c' , d') = ℂ [ c , c' ] × 𝔻 [ d , d' ]
|
2018-01-24 15:38:28 +00:00
|
|
|
|
:𝟙: : {o : :Object:} → :Arrow: o o
|
2018-02-05 15:35:33 +00:00
|
|
|
|
:𝟙: = 𝟙 ℂ , 𝟙 𝔻
|
2018-01-24 15:38:28 +00:00
|
|
|
|
_:⊕:_ :
|
|
|
|
|
{a b c : :Object:} →
|
|
|
|
|
:Arrow: b c →
|
|
|
|
|
:Arrow: a b →
|
|
|
|
|
:Arrow: a c
|
2018-01-30 18:19:16 +00:00
|
|
|
|
_:⊕:_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) → ℂ [ bc∈C ∘ ab∈C ] , 𝔻 [ bc∈D ∘ ab∈D ]}
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-02-05 15:35:33 +00:00
|
|
|
|
:rawProduct: : RawCategory ℓ ℓ'
|
|
|
|
|
RawCategory.Object :rawProduct: = :Object:
|
|
|
|
|
RawCategory.Arrow :rawProduct: = :Arrow:
|
|
|
|
|
RawCategory.𝟙 :rawProduct: = :𝟙:
|
|
|
|
|
RawCategory._∘_ :rawProduct: = _:⊕:_
|
2018-02-23 09:34:37 +00:00
|
|
|
|
open RawCategory :rawProduct:
|
2018-02-05 15:35:33 +00:00
|
|
|
|
|
2018-02-22 14:31:54 +00:00
|
|
|
|
module C = Category ℂ
|
|
|
|
|
module D = Category 𝔻
|
2018-02-23 11:33:20 +00:00
|
|
|
|
open import Cubical.Sigma
|
|
|
|
|
issSet : {A B : RawCategory.Object :rawProduct:} → isSet (Arrow A B)
|
2018-02-23 11:51:44 +00:00
|
|
|
|
issSet = setSig {sA = C.arrowsAreSets} {sB = λ x → D.arrowsAreSets}
|
2018-02-23 09:34:37 +00:00
|
|
|
|
ident' : IsIdentity :𝟙:
|
|
|
|
|
ident'
|
2018-02-23 11:49:41 +00:00
|
|
|
|
= Σ≡ (fst C.isIdentity) (fst D.isIdentity)
|
|
|
|
|
, Σ≡ (snd C.isIdentity) (snd D.isIdentity)
|
2018-02-23 09:34:37 +00:00
|
|
|
|
postulate univalent : Univalence.Univalent :rawProduct: ident'
|
2018-01-24 15:38:28 +00:00
|
|
|
|
instance
|
2018-02-05 15:35:33 +00:00
|
|
|
|
:isCategory: : IsCategory :rawProduct:
|
2018-02-23 11:43:49 +00:00
|
|
|
|
IsCategory.isAssociative :isCategory: = Σ≡ C.isAssociative D.isAssociative
|
2018-02-23 11:49:41 +00:00
|
|
|
|
IsCategory.isIdentity :isCategory: = ident'
|
2018-02-23 11:51:44 +00:00
|
|
|
|
IsCategory.arrowsAreSets :isCategory: = issSet
|
2018-02-23 09:34:37 +00:00
|
|
|
|
IsCategory.univalent :isCategory: = univalent
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
|
|
|
|
:product: : Category ℓ ℓ'
|
2018-02-22 14:31:54 +00:00
|
|
|
|
Category.raw :product: = :rawProduct:
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-01-31 13:39:54 +00:00
|
|
|
|
proj₁ : Catt [ :product: , ℂ ]
|
2018-02-22 14:31:54 +00:00
|
|
|
|
proj₁ = record
|
|
|
|
|
{ raw = record { func* = fst ; func→ = fst }
|
2018-02-23 11:53:35 +00:00
|
|
|
|
; isFunctor = record { isIdentity = refl ; isDistributive = refl }
|
2018-02-22 14:31:54 +00:00
|
|
|
|
}
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-01-31 13:39:54 +00:00
|
|
|
|
proj₂ : Catt [ :product: , 𝔻 ]
|
2018-02-22 14:31:54 +00:00
|
|
|
|
proj₂ = record
|
|
|
|
|
{ raw = record { func* = snd ; func→ = snd }
|
2018-02-23 11:53:35 +00:00
|
|
|
|
; isFunctor = record { isIdentity = refl ; isDistributive = refl }
|
2018-02-22 14:31:54 +00:00
|
|
|
|
}
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-01-31 13:39:54 +00:00
|
|
|
|
module _ {X : Object Catt} (x₁ : Catt [ X , ℂ ]) (x₂ : Catt [ X , 𝔻 ]) where
|
2018-02-23 11:29:10 +00:00
|
|
|
|
x : Functor X :product:
|
|
|
|
|
x = record
|
|
|
|
|
{ raw = record
|
|
|
|
|
{ func* = λ x → x₁ .func* x , x₂ .func* x
|
|
|
|
|
; func→ = λ x → func→ x₁ x , func→ x₂ x
|
|
|
|
|
}
|
|
|
|
|
; isFunctor = record
|
2018-02-23 11:49:41 +00:00
|
|
|
|
{ isIdentity = Σ≡ x₁.isIdentity x₂.isIdentity
|
2018-02-23 11:53:35 +00:00
|
|
|
|
; isDistributive = Σ≡ x₁.isDistributive x₂.isDistributive
|
2018-02-23 11:29:10 +00:00
|
|
|
|
}
|
|
|
|
|
}
|
|
|
|
|
where
|
|
|
|
|
open module x₁ = Functor x₁
|
|
|
|
|
open module x₂ = Functor x₂
|
|
|
|
|
|
|
|
|
|
isUniqL : Catt [ proj₁ ∘ x ] ≡ x₁
|
|
|
|
|
isUniqL = Functor≡ eq* eq→
|
|
|
|
|
where
|
|
|
|
|
eq* : (Catt [ proj₁ ∘ x ]) .func* ≡ x₁ .func*
|
|
|
|
|
eq* = refl
|
|
|
|
|
eq→ : (λ i → {A : Object X} {B : Object X} → X [ A , B ] → ℂ [ eq* i A , eq* i B ])
|
|
|
|
|
[ (Catt [ proj₁ ∘ x ]) .func→ ≡ x₁ .func→ ]
|
|
|
|
|
eq→ = refl
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-02-23 11:29:10 +00:00
|
|
|
|
isUniqR : Catt [ proj₂ ∘ x ] ≡ x₂
|
|
|
|
|
isUniqR = Functor≡ refl refl
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-01-30 18:19:16 +00:00
|
|
|
|
isUniq : Catt [ proj₁ ∘ x ] ≡ x₁ × Catt [ proj₂ ∘ x ] ≡ x₂
|
2018-01-24 15:38:28 +00:00
|
|
|
|
isUniq = isUniqL , isUniqR
|
|
|
|
|
|
2018-01-30 18:19:16 +00:00
|
|
|
|
uniq : ∃![ x ] (Catt [ proj₁ ∘ x ] ≡ x₁ × Catt [ proj₂ ∘ x ] ≡ x₂)
|
2018-01-24 15:38:28 +00:00
|
|
|
|
uniq = x , isUniq
|
|
|
|
|
|
2018-01-25 11:44:47 +00:00
|
|
|
|
instance
|
2018-02-23 09:34:37 +00:00
|
|
|
|
isProduct : IsProduct Catt proj₁ proj₂
|
2018-01-25 11:44:47 +00:00
|
|
|
|
isProduct = uniq
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-02-23 09:34:37 +00:00
|
|
|
|
product : Product {ℂ = Catt} ℂ 𝔻
|
2018-01-24 15:38:28 +00:00
|
|
|
|
product = record
|
|
|
|
|
{ obj = :product:
|
|
|
|
|
; proj₁ = proj₁
|
|
|
|
|
; proj₂ = proj₂
|
2018-01-21 13:31:37 +00:00
|
|
|
|
}
|
|
|
|
|
|
2018-02-23 09:34:37 +00:00
|
|
|
|
module _ {ℓ ℓ' : Level} (unprovable : IsCategory (RawCat ℓ ℓ')) where
|
|
|
|
|
Catt = Cat ℓ ℓ' unprovable
|
2018-01-24 15:38:28 +00:00
|
|
|
|
instance
|
2018-02-23 09:34:37 +00:00
|
|
|
|
hasProducts : HasProducts Catt
|
|
|
|
|
hasProducts = record { product = product unprovable }
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-01-24 15:38:28 +00:00
|
|
|
|
-- Basically proves that `Cat ℓ ℓ` is cartesian closed.
|
2018-02-23 09:34:37 +00:00
|
|
|
|
module _ (ℓ : Level) (unprovable : IsCategory (RawCat ℓ ℓ)) where
|
2018-01-24 15:38:28 +00:00
|
|
|
|
private
|
2018-01-25 11:01:37 +00:00
|
|
|
|
open Data.Product
|
2018-01-24 15:38:28 +00:00
|
|
|
|
open import Cat.Categories.Fun
|
2018-01-25 11:01:37 +00:00
|
|
|
|
|
|
|
|
|
Catℓ : Category (lsuc (ℓ ⊔ ℓ)) (ℓ ⊔ ℓ)
|
2018-02-23 09:34:37 +00:00
|
|
|
|
Catℓ = Cat ℓ ℓ unprovable
|
2018-01-24 15:38:28 +00:00
|
|
|
|
module _ (ℂ 𝔻 : Category ℓ ℓ) where
|
2018-02-23 16:33:09 +00:00
|
|
|
|
open Fun ℂ 𝔻 renaming (identity to idN)
|
2018-01-24 15:38:28 +00:00
|
|
|
|
private
|
2018-02-23 09:34:37 +00:00
|
|
|
|
:obj: : Object Catℓ
|
2018-02-23 16:33:09 +00:00
|
|
|
|
:obj: = Fun
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-02-05 15:35:33 +00:00
|
|
|
|
:func*: : Functor ℂ 𝔻 × Object ℂ → Object 𝔻
|
2018-02-22 14:31:54 +00:00
|
|
|
|
:func*: (F , A) = func* F A
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-02-05 15:35:33 +00:00
|
|
|
|
module _ {dom cod : Functor ℂ 𝔻 × Object ℂ} where
|
2018-01-24 15:38:28 +00:00
|
|
|
|
private
|
|
|
|
|
F : Functor ℂ 𝔻
|
|
|
|
|
F = proj₁ dom
|
2018-02-05 15:35:33 +00:00
|
|
|
|
A : Object ℂ
|
2018-01-24 15:38:28 +00:00
|
|
|
|
A = proj₂ dom
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-01-24 15:38:28 +00:00
|
|
|
|
G : Functor ℂ 𝔻
|
|
|
|
|
G = proj₁ cod
|
2018-02-05 15:35:33 +00:00
|
|
|
|
B : Object ℂ
|
2018-01-24 15:38:28 +00:00
|
|
|
|
B = proj₂ cod
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-02-05 15:35:33 +00:00
|
|
|
|
:func→: : (pobj : NaturalTransformation F G × ℂ [ A , B ])
|
2018-02-22 14:31:54 +00:00
|
|
|
|
→ 𝔻 [ func* F A , func* G B ]
|
2018-01-24 15:38:28 +00:00
|
|
|
|
:func→: ((θ , θNat) , f) = result
|
|
|
|
|
where
|
2018-02-22 14:31:54 +00:00
|
|
|
|
θA : 𝔻 [ func* F A , func* G A ]
|
2018-01-24 15:38:28 +00:00
|
|
|
|
θA = θ A
|
2018-02-22 14:31:54 +00:00
|
|
|
|
θB : 𝔻 [ func* F B , func* G B ]
|
2018-01-24 15:38:28 +00:00
|
|
|
|
θB = θ B
|
2018-02-22 14:31:54 +00:00
|
|
|
|
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 ]
|
2018-01-30 18:19:16 +00:00
|
|
|
|
l = 𝔻 [ θB ∘ F→f ]
|
2018-02-22 14:31:54 +00:00
|
|
|
|
r : 𝔻 [ func* F A , func* G B ]
|
2018-01-30 18:19:16 +00:00
|
|
|
|
r = 𝔻 [ G→f ∘ θA ]
|
2018-01-24 15:38:28 +00:00
|
|
|
|
-- There are two choices at this point,
|
|
|
|
|
-- but I suppose the whole point is that
|
|
|
|
|
-- by `θNat f` we have `l ≡ r`
|
2018-01-30 18:19:16 +00:00
|
|
|
|
-- lem : 𝔻 [ θ B ∘ F .func→ f ] ≡ 𝔻 [ G .func→ f ∘ θ A ]
|
2018-01-24 15:38:28 +00:00
|
|
|
|
-- lem = θNat f
|
2018-02-22 14:31:54 +00:00
|
|
|
|
result : 𝔻 [ func* F A , func* G B ]
|
2018-01-24 15:38:28 +00:00
|
|
|
|
result = l
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-02-23 09:34:37 +00:00
|
|
|
|
_×p_ = product unprovable
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-02-05 15:35:33 +00:00
|
|
|
|
module _ {c : Functor ℂ 𝔻 × Object ℂ} where
|
2018-01-24 15:38:28 +00:00
|
|
|
|
private
|
|
|
|
|
F : Functor ℂ 𝔻
|
|
|
|
|
F = proj₁ c
|
2018-02-05 15:35:33 +00:00
|
|
|
|
C : Object ℂ
|
2018-01-24 15:38:28 +00:00
|
|
|
|
C = proj₂ c
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-01-24 15:38:28 +00:00
|
|
|
|
-- NaturalTransformation F G × ℂ .Arrow A B
|
2018-01-25 11:01:37 +00:00
|
|
|
|
-- :ident: : :func→: {c} {c} (identityNat F , ℂ .𝟙) ≡ 𝔻 .𝟙
|
2018-02-23 11:49:41 +00:00
|
|
|
|
-- :ident: = trans (proj₂ 𝔻.isIdentity) (F .isIdentity)
|
2018-01-24 15:38:28 +00:00
|
|
|
|
-- where
|
|
|
|
|
-- open module 𝔻 = IsCategory (𝔻 .isCategory)
|
2018-01-25 11:01:37 +00:00
|
|
|
|
-- Unfortunately the equational version has some ambigous arguments.
|
2018-02-23 16:33:09 +00:00
|
|
|
|
|
|
|
|
|
:ident: : :func→: {c} {c} (NT.identity F , 𝟙 ℂ {A = proj₂ c}) ≡ 𝟙 𝔻
|
2018-01-25 11:01:37 +00:00
|
|
|
|
:ident: = begin
|
2018-02-05 15:35:33 +00:00
|
|
|
|
:func→: {c} {c} (𝟙 (Product.obj (:obj: ×p ℂ)) {c}) ≡⟨⟩
|
2018-02-23 16:33:09 +00:00
|
|
|
|
:func→: {c} {c} (idN F , 𝟙 ℂ) ≡⟨⟩
|
2018-02-22 14:31:54 +00:00
|
|
|
|
𝔻 [ identityTrans F C ∘ func→ F (𝟙 ℂ)] ≡⟨⟩
|
2018-02-23 11:49:41 +00:00
|
|
|
|
𝔻 [ 𝟙 𝔻 ∘ func→ F (𝟙 ℂ)] ≡⟨ proj₂ 𝔻.isIdentity ⟩
|
|
|
|
|
func→ F (𝟙 ℂ) ≡⟨ F.isIdentity ⟩
|
2018-02-05 15:35:33 +00:00
|
|
|
|
𝟙 𝔻 ∎
|
2018-01-25 11:01:37 +00:00
|
|
|
|
where
|
2018-02-22 14:31:54 +00:00
|
|
|
|
open module 𝔻 = Category 𝔻
|
2018-02-23 11:29:10 +00:00
|
|
|
|
open module F = Functor F
|
2018-01-30 15:23:36 +00:00
|
|
|
|
|
2018-02-23 09:44:23 +00:00
|
|
|
|
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._∘_ (Product.obj (:obj: ×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
|
|
|
|
|
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
|
|
|
|
|
|
2018-02-23 11:53:35 +00:00
|
|
|
|
:isDistributive: :
|
2018-02-23 09:44:23 +00:00
|
|
|
|
𝔻 [ 𝔻 [ η C ∘ θ C ] ∘ func→ F ( ℂ [ g ∘ f ] ) ]
|
|
|
|
|
≡ 𝔻 [ 𝔻 [ η C ∘ func→ G g ] ∘ 𝔻 [ θ B ∘ func→ F f ] ]
|
2018-02-23 11:53:35 +00:00
|
|
|
|
:isDistributive: = begin
|
2018-02-23 09:44:23 +00:00
|
|
|
|
𝔻 [ (ηθ C) ∘ func→ F (ℂ [ g ∘ f ]) ]
|
|
|
|
|
≡⟨ ηθNat (ℂ [ g ∘ f ]) ⟩
|
|
|
|
|
𝔻 [ func→ H (ℂ [ g ∘ f ]) ∘ (ηθ A) ]
|
2018-02-23 11:53:35 +00:00
|
|
|
|
≡⟨ cong (λ φ → 𝔻 [ φ ∘ ηθ A ]) (H.isDistributive) ⟩
|
2018-02-23 09:44:23 +00:00
|
|
|
|
𝔻 [ 𝔻 [ func→ H g ∘ func→ H f ] ∘ (ηθ A) ]
|
2018-02-23 11:43:49 +00:00
|
|
|
|
≡⟨ sym isAssociative ⟩
|
2018-02-23 09:44:23 +00:00
|
|
|
|
𝔻 [ func→ H g ∘ 𝔻 [ func→ H f ∘ ηθ A ] ]
|
2018-02-23 11:43:49 +00:00
|
|
|
|
≡⟨ cong (λ φ → 𝔻 [ func→ H g ∘ φ ]) isAssociative ⟩
|
2018-02-23 09:44:23 +00:00
|
|
|
|
𝔻 [ 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 ] ]
|
2018-02-23 11:43:49 +00:00
|
|
|
|
≡⟨ cong (λ φ → 𝔻 [ func→ H g ∘ φ ]) (sym isAssociative) ⟩
|
2018-02-23 09:44:23 +00:00
|
|
|
|
𝔻 [ func→ H g ∘ 𝔻 [ η B ∘ 𝔻 [ func→ G f ∘ θ A ] ] ]
|
2018-02-23 11:43:49 +00:00
|
|
|
|
≡⟨ isAssociative ⟩
|
2018-02-23 09:44:23 +00:00
|
|
|
|
𝔻 [ 𝔻 [ 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 𝔻
|
2018-02-23 11:29:10 +00:00
|
|
|
|
module H = Functor H
|
2018-02-23 09:44:23 +00:00
|
|
|
|
|
|
|
|
|
:eval: : Functor ((:obj: ×p ℂ) .Product.obj) 𝔻
|
|
|
|
|
:eval: = record
|
|
|
|
|
{ raw = record
|
|
|
|
|
{ func* = :func*:
|
|
|
|
|
; func→ = λ {dom} {cod} → :func→: {dom} {cod}
|
|
|
|
|
}
|
|
|
|
|
; isFunctor = record
|
2018-02-23 11:49:41 +00:00
|
|
|
|
{ isIdentity = λ {o} → :ident: {o}
|
2018-02-23 11:53:35 +00:00
|
|
|
|
; 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 ℂ) .Product.obj) 𝔻) where
|
|
|
|
|
open HasProducts (hasProducts {ℓ} {ℓ} unprovable) renaming (_|×|_ to parallelProduct)
|
|
|
|
|
|
|
|
|
|
postulate
|
|
|
|
|
transpose : Functor 𝔸 :obj:
|
|
|
|
|
eq : Catℓ [ :eval: ∘ (parallelProduct transpose (𝟙 Catℓ {A = ℂ})) ] ≡ 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
|
|
|
|
|
|
|
|
|
|
postulate :isExponential: : IsExponential Catℓ ℂ 𝔻 :obj: :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 = :obj:
|
|
|
|
|
; eval = :eval:
|
|
|
|
|
; isExponential = :isExponential:
|
|
|
|
|
}
|
|
|
|
|
|
|
|
|
|
hasExponentials : HasExponentials Catℓ
|
|
|
|
|
hasExponentials = record { exponent = :exponent: }
|