Clean-up yoneda embedding
This commit is contained in:
parent
1bf565b87a
commit
7fbca1aeeb
|
@ -5,23 +5,18 @@ module Cat.Category.Yoneda where
|
||||||
open import Agda.Primitive
|
open import Agda.Primitive
|
||||||
open import Data.Product
|
open import Data.Product
|
||||||
open import Cubical
|
open import Cubical
|
||||||
|
open import Cubical.NType.Properties
|
||||||
|
|
||||||
open import Cat.Category
|
open import Cat.Category
|
||||||
open import Cat.Category.Functor
|
open import Cat.Category.Functor
|
||||||
open import Cat.Equality
|
open import Cat.Equality
|
||||||
open Equality.Data.Product
|
|
||||||
|
|
||||||
-- TODO: We want to avoid defining the yoneda embedding going through the
|
open import Cat.Categories.Fun
|
||||||
-- category of categories (since it doesn't exist).
|
open import Cat.Categories.Sets
|
||||||
open import Cat.Categories.Cat using (RawCat)
|
open import Cat.Categories.Cat
|
||||||
|
|
||||||
module _ {ℓ : Level} {ℂ : Category ℓ ℓ} where
|
module _ {ℓ : Level} {ℂ : Category ℓ ℓ} where
|
||||||
private
|
private
|
||||||
open import Cat.Categories.Fun
|
|
||||||
open import Cat.Categories.Sets
|
|
||||||
module Cat = Cat.Categories.Cat
|
|
||||||
open import Cat.Category.Exponential
|
|
||||||
open Functor
|
|
||||||
𝓢 = Sets ℓ
|
𝓢 = Sets ℓ
|
||||||
open Fun (opposite ℂ) 𝓢
|
open Fun (opposite ℂ) 𝓢
|
||||||
prshf = presheaf ℂ
|
prshf = presheaf ℂ
|
||||||
|
@ -34,33 +29,31 @@ module _ {ℓ : Level} {ℂ : Category ℓ ℓ} where
|
||||||
--
|
--
|
||||||
-- In stead we'll use an ad-hoc definition -- which is definitionally
|
-- In stead we'll use an ad-hoc definition -- which is definitionally
|
||||||
-- equivalent to that other one.
|
-- equivalent to that other one.
|
||||||
_⇑_ = Cat.CatExponential.prodObj
|
_⇑_ = CatExponential.prodObj
|
||||||
|
|
||||||
module _ {A B : ℂ.Object} (f : ℂ [ A , B ]) where
|
module _ {A B : ℂ.Object} (f : ℂ [ A , B ]) where
|
||||||
:func→: : NaturalTransformation (prshf A) (prshf B)
|
:func→: : NaturalTransformation (prshf A) (prshf B)
|
||||||
:func→: = (λ C x → ℂ [ f ∘ x ]) , λ f₁ → funExt λ _ → ℂ.isAssociative
|
:func→: = (λ C x → ℂ [ f ∘ x ]) , λ f₁ → funExt λ _ → ℂ.isAssociative
|
||||||
|
|
||||||
module _ {c : Category.Object ℂ} where
|
rawYoneda : RawFunctor ℂ Fun
|
||||||
eqTrans : (λ _ → Transformation (prshf c) (prshf c))
|
RawFunctor.func* rawYoneda = prshf
|
||||||
[ (λ _ x → ℂ [ ℂ.𝟙 ∘ x ]) ≡ identityTrans (prshf c) ]
|
RawFunctor.func→ rawYoneda = :func→:
|
||||||
eqTrans = funExt λ x → funExt λ x → ℂ.isIdentity .proj₂
|
open RawFunctor rawYoneda
|
||||||
|
|
||||||
open import Cubical.NType.Properties
|
isIdentity : IsIdentity
|
||||||
open import Cat.Categories.Fun
|
isIdentity {c} = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq
|
||||||
:ident: : :func→: (ℂ.𝟙 {c}) ≡ Category.𝟙 Fun {A = prshf c}
|
|
||||||
:ident: = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq
|
|
||||||
where
|
where
|
||||||
eq : (λ C x → ℂ [ ℂ.𝟙 ∘ x ]) ≡ identityTrans (prshf c)
|
eq : (λ C x → ℂ [ ℂ.𝟙 ∘ x ]) ≡ identityTrans (prshf c)
|
||||||
eq = funExt λ A → funExt λ B → proj₂ ℂ.isIdentity
|
eq = funExt λ A → funExt λ B → proj₂ ℂ.isIdentity
|
||||||
|
|
||||||
|
isDistributive : IsDistributive
|
||||||
|
isDistributive = {!!}
|
||||||
|
|
||||||
|
instance
|
||||||
|
isFunctor : IsFunctor ℂ Fun rawYoneda
|
||||||
|
IsFunctor.isIdentity isFunctor = isIdentity
|
||||||
|
IsFunctor.isDistributive isFunctor = isDistributive
|
||||||
|
|
||||||
yoneda : Functor ℂ Fun
|
yoneda : Functor ℂ Fun
|
||||||
yoneda = record
|
Functor.raw yoneda = rawYoneda
|
||||||
{ raw = record
|
Functor.isFunctor yoneda = isFunctor
|
||||||
{ func* = prshf
|
|
||||||
; func→ = :func→:
|
|
||||||
}
|
|
||||||
; isFunctor = record
|
|
||||||
{ isIdentity = :ident:
|
|
||||||
; isDistributive = {!!}
|
|
||||||
}
|
|
||||||
}
|
|
||||||
|
|
Loading…
Reference in a new issue