Formatting in yoneda

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-20 16:43:48 +01:00
parent b6a9befd9c
commit 629115661b

View file

@ -12,51 +12,52 @@ open import Cat.Category.Functor
open import Cat.Equality
open import Cat.Categories.Fun
open import Cat.Categories.Sets
open import Cat.Categories.Cat
open import Cat.Categories.Sets hiding (presheaf)
-- There is no (small) category of categories. So we won't use _⇑_ from
-- `HasExponential`
--
-- open HasExponentials (Cat.hasExponentials unprovable) using (_⇑_)
--
-- In stead we'll use an ad-hoc definition -- which is definitionally equivalent
-- to that other one - even without mentioning the category of categories.
_⇑_ : { : Level} Category Category Category
_⇑_ = Fun.Fun
module _ { : Level} { : Category } where
private
𝓢 = Sets
open Fun (opposite ) 𝓢
prshf = presheaf
presheaf = Cat.Categories.Sets.presheaf
module = Category
-- There is no (small) category of categories. So we won't use _⇑_ from
-- `HasExponential`
--
-- open HasExponentials (Cat.hasExponentials unprovable) using (_⇑_)
--
-- In stead we'll use an ad-hoc definition -- which is definitionally
-- equivalent to that other one.
_⇑_ = CatExponential.object
module _ {A B : .Object} (f : [ A , B ]) where
fmap : Transformation (prshf A) (prshf B)
fmap : Transformation (presheaf A) (presheaf B)
fmap C x = [ f x ]
fmapNatural : Natural (prshf A) (prshf B) fmap
fmapNatural : Natural (presheaf A) (presheaf B) fmap
fmapNatural g = funExt λ _ .isAssociative
fmapNT : NaturalTransformation (prshf A) (prshf B)
fmapNT : NaturalTransformation (presheaf A) (presheaf B)
fmapNT = fmap , fmapNatural
rawYoneda : RawFunctor Fun
RawFunctor.omap rawYoneda = prshf
RawFunctor.omap rawYoneda = presheaf
RawFunctor.fmap rawYoneda = fmapNT
open RawFunctor rawYoneda hiding (fmap)
isIdentity : IsIdentity
isIdentity {c} = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq
isIdentity {c} = lemSig (naturalIsProp {F = presheaf c} {presheaf c}) _ _ eq
where
eq : (λ C x [ .𝟙 x ]) identityTrans (prshf c)
eq : (λ C x [ .𝟙 x ]) identityTrans (presheaf c)
eq = funExt λ A funExt λ B proj₂ .isIdentity
isDistributive : IsDistributive
isDistributive {A} {B} {C} {f = f} {g}
= lemSig (propIsNatural (prshf A) (prshf C)) _ _ eq
= lemSig (propIsNatural (presheaf A) (presheaf C)) _ _ eq
where
T[_∘_]' = T[_∘_] {F = prshf A} {prshf B} {prshf C}
T[_∘_]' = T[_∘_] {F = presheaf A} {presheaf B} {presheaf C}
eqq : (X : .Object) (x : [ X , A ])
fmap ( [ g f ]) X x T[ fmap g fmap f ]' X x
eqq X x = begin
@ -76,5 +77,5 @@ module _ { : Level} { : Category } where
IsFunctor.isDistributive isFunctor = isDistributive
yoneda : Functor Fun
Functor.raw yoneda = rawYoneda
Functor.raw yoneda = rawYoneda
Functor.isFunctor yoneda = isFunctor