2018-01-24 15:38:28 +00:00
|
|
|
|
{-# OPTIONS --allow-unsolved-metas --cubical #-}
|
2018-01-15 15:13:23 +00:00
|
|
|
|
|
2018-02-25 14:28:42 +00:00
|
|
|
|
module Cat.Category.Yoneda where
|
2018-01-15 15:13:23 +00:00
|
|
|
|
|
2018-01-21 14:01:01 +00:00
|
|
|
|
open import Agda.Primitive
|
|
|
|
|
open import Data.Product
|
2018-01-30 10:19:48 +00:00
|
|
|
|
open import Cubical
|
2018-01-21 14:01:01 +00:00
|
|
|
|
|
2018-01-15 15:13:23 +00:00
|
|
|
|
open import Cat.Category
|
2018-02-05 13:59:53 +00:00
|
|
|
|
open import Cat.Category.Functor
|
2018-01-30 21:41:18 +00:00
|
|
|
|
open import Cat.Equality
|
|
|
|
|
open Equality.Data.Product
|
2018-01-15 15:13:23 +00:00
|
|
|
|
|
2018-02-23 09:53:11 +00:00
|
|
|
|
-- TODO: We want to avoid defining the yoneda embedding going through the
|
|
|
|
|
-- category of categories (since it doesn't exist).
|
|
|
|
|
open import Cat.Categories.Cat using (RawCat)
|
|
|
|
|
|
|
|
|
|
module _ {ℓ : Level} {ℂ : Category ℓ ℓ} (unprovable : IsCategory (RawCat ℓ ℓ)) where
|
|
|
|
|
open import Cat.Categories.Fun
|
|
|
|
|
open import Cat.Categories.Sets
|
|
|
|
|
module Cat = Cat.Categories.Cat
|
|
|
|
|
open import Cat.Category.Exponential
|
|
|
|
|
open Functor
|
|
|
|
|
𝓢 = Sets ℓ
|
2018-02-25 14:23:33 +00:00
|
|
|
|
open Fun (opposite ℂ) 𝓢
|
2018-02-23 09:53:11 +00:00
|
|
|
|
private
|
|
|
|
|
Catℓ : Category _ _
|
|
|
|
|
Catℓ = record { raw = RawCat ℓ ℓ ; isCategory = unprovable}
|
|
|
|
|
prshf = presheaf {ℂ = ℂ}
|
|
|
|
|
module ℂ = Category ℂ
|
|
|
|
|
|
|
|
|
|
_⇑_ : (A B : Category.Object Catℓ) → Category.Object Catℓ
|
|
|
|
|
A ⇑ B = (exponent A B) .obj
|
|
|
|
|
where
|
|
|
|
|
open HasExponentials (Cat.hasExponentials ℓ unprovable)
|
|
|
|
|
|
|
|
|
|
module _ {A B : ℂ.Object} (f : ℂ [ A , B ]) where
|
|
|
|
|
:func→: : NaturalTransformation (prshf A) (prshf B)
|
2018-02-23 11:43:49 +00:00
|
|
|
|
:func→: = (λ C x → ℂ [ f ∘ x ]) , λ f₁ → funExt λ _ → ℂ.isAssociative
|
2018-02-23 09:53:11 +00:00
|
|
|
|
|
|
|
|
|
module _ {c : Category.Object ℂ} where
|
|
|
|
|
eqTrans : (λ _ → Transformation (prshf c) (prshf c))
|
|
|
|
|
[ (λ _ x → ℂ [ ℂ.𝟙 ∘ x ]) ≡ identityTrans (prshf c) ]
|
2018-02-23 11:49:41 +00:00
|
|
|
|
eqTrans = funExt λ x → funExt λ x → ℂ.isIdentity .proj₂
|
2018-02-23 09:53:11 +00:00
|
|
|
|
|
2018-02-23 10:23:21 +00:00
|
|
|
|
open import Cubical.NType.Properties
|
|
|
|
|
open import Cat.Categories.Fun
|
|
|
|
|
:ident: : :func→: (ℂ.𝟙 {c}) ≡ Category.𝟙 Fun {A = prshf c}
|
|
|
|
|
:ident: = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq
|
|
|
|
|
where
|
|
|
|
|
eq : (λ C x → ℂ [ ℂ.𝟙 ∘ x ]) ≡ identityTrans (prshf c)
|
2018-02-23 11:49:41 +00:00
|
|
|
|
eq = funExt λ A → funExt λ B → proj₂ ℂ.isIdentity
|
2018-02-23 09:53:11 +00:00
|
|
|
|
|
2018-02-23 16:33:09 +00:00
|
|
|
|
yoneda : Functor ℂ Fun
|
2018-02-23 09:53:11 +00:00
|
|
|
|
yoneda = record
|
|
|
|
|
{ raw = record
|
|
|
|
|
{ func* = prshf
|
|
|
|
|
; func→ = :func→:
|
|
|
|
|
}
|
|
|
|
|
; isFunctor = record
|
2018-02-23 11:49:41 +00:00
|
|
|
|
{ isIdentity = :ident:
|
2018-02-23 11:53:35 +00:00
|
|
|
|
; isDistributive = {!!}
|
2018-02-23 09:53:11 +00:00
|
|
|
|
}
|
|
|
|
|
}
|