cat/src/Cat/Category/Yoneda.agda

65 lines
2.1 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.

{-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Category.Yoneda where
open import Agda.Primitive
open import Data.Product
open import Cubical
open import Cat.Category
open import Cat.Category.Functor
open import Cat.Equality
open Equality.Data.Product
-- 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
open Fun (opposite ) 𝓢
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)
:func→: = (λ C x [ f x ]) , λ f₁ funExt λ _ .isAssociative
module _ {c : Category.Object } where
eqTrans : (λ _ Transformation (prshf c) (prshf c))
[ (λ _ x [ .𝟙 x ]) identityTrans (prshf c) ]
eqTrans = funExt λ x funExt λ x .isIdentity .proj₂
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)
eq = funExt λ A funExt λ B proj₂ .isIdentity
yoneda : Functor Fun
yoneda = record
{ raw = record
{ func* = prshf
; func→ = :func→:
}
; isFunctor = record
{ isIdentity = :ident:
; isDistributive = {!!}
}
}