Move properties about natural transformations to that module

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-23 15:20:26 +01:00
parent ef688202a2
commit d3864dbae5
8 changed files with 171 additions and 179 deletions

View file

@ -9,7 +9,7 @@ open import Cat.Category
open import Cat.Category.Functor open import Cat.Category.Functor
open import Cat.Category.Product open import Cat.Category.Product
open import Cat.Category.Exponential hiding (_×_ ; product) open import Cat.Category.Exponential hiding (_×_ ; product)
open import Cat.Category.NaturalTransformation import Cat.Category.NaturalTransformation
open import Cat.Categories.Fun open import Cat.Categories.Fun
-- The category of categories -- The category of categories
@ -155,6 +155,9 @@ module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
-- | The category of categories have expoentntials - and because it has products -- | The category of categories have expoentntials - and because it has products
-- it is therefory also cartesian closed. -- it is therefory also cartesian closed.
module CatExponential { : Level} ( 𝔻 : Category ) where module CatExponential { : Level} ( 𝔻 : Category ) where
open Cat.Category.NaturalTransformation 𝔻
renaming (identity to identityNT)
using ()
private private
module = Category module = Category
module 𝔻 = Category 𝔻 module 𝔻 = Category 𝔻
@ -189,7 +192,7 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
module _ {c : Functor 𝔻 × .Object} where module _ {c : Functor 𝔻 × .Object} where
open Σ c renaming (proj₁ to F ; proj₂ to C) open Σ c renaming (proj₁ to F ; proj₂ to C)
ident : fmap {c} {c} (NT.identity F , .𝟙 {A = snd c}) 𝔻.𝟙 ident : fmap {c} {c} (identityNT F , .𝟙 {A = snd c}) 𝔻.𝟙
ident = begin ident = begin
fmap {c} {c} (Category.𝟙 (object ) {c}) ≡⟨⟩ fmap {c} {c} (Category.𝟙 (object ) {c}) ≡⟨⟩
fmap {c} {c} (idN F , .𝟙) ≡⟨⟩ fmap {c} {c} (idN F , .𝟙) ≡⟨⟩

View file

@ -5,62 +5,26 @@ open import Cat.Prelude
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor open import Cat.Category.Functor
open import Cat.Category.NaturalTransformation
module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : Category d d') where module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : Category d d') where
module NT = NaturalTransformation 𝔻 import Cat.Category.NaturalTransformation 𝔻
open NT public as NaturalTransformation
open NaturalTransformation public hiding (module Properties)
open NaturalTransformation.Properties
private private
module = Category module = Category
module 𝔻 = Category 𝔻 module 𝔻 = Category 𝔻
private
module _ {A B C D : Functor 𝔻} {θNT : NaturalTransformation A B}
{ηNT : NaturalTransformation B C} {ζNT : NaturalTransformation C D} where
open Σ θNT renaming (proj₁ to θ ; proj₂ to θNat)
open Σ ηNT renaming (proj₁ to η ; proj₂ to ηNat)
open Σ ζNT renaming (proj₁ to ζ ; proj₂ to ζNat)
private
L : NaturalTransformation A D
L = (NT[_∘_] {A} {C} {D} ζNT (NT[_∘_] {A} {B} {C} ηNT θNT))
R : NaturalTransformation A D
R = (NT[_∘_] {A} {B} {D} (NT[_∘_] {B} {C} {D} ζNT ηNT) θNT)
_g⊕f_ = NT[_∘_] {A} {B} {C}
_h⊕g_ = NT[_∘_] {B} {C} {D}
isAssociative : L R
isAssociative = lemSig (naturalIsProp {F = A} {D})
L R (funExt (λ x 𝔻.isAssociative))
private -- Functor categories. Objects are functors, arrows are natural transformations.
module _ {A B : Functor 𝔻} {f : NaturalTransformation A B} where raw : RawCategory (c c' d d') (c c' d')
allNatural = naturalIsProp {F = A} {B} RawCategory.Object raw = Functor 𝔻
f' = proj₁ f RawCategory.Arrow raw = NaturalTransformation
eq-r : C (𝔻 [ f' C identityTrans A C ]) f' C RawCategory.𝟙 raw {F} = identity F
eq-r C = begin RawCategory._∘_ raw {F} {G} {H} = NT[_∘_] {F} {G} {H}
𝔻 [ f' C identityTrans A C ] ≡⟨⟩
𝔻 [ f' C 𝔻.𝟙 ] ≡⟨ 𝔻.rightIdentity open RawCategory raw
f' C open Univalence (λ {A} {B} {f} isIdentity {F = A} {B} {f})
eq-l : C (𝔻 [ identityTrans B C f' C ]) f' C
eq-l C = 𝔻.leftIdentity
ident-r : (NT[_∘_] {A} {A} {B} f (NT.identity A)) f
ident-r = lemSig allNatural _ _ (funExt eq-r)
ident-l : (NT[_∘_] {A} {B} {B} (NT.identity B) f) f
ident-l = lemSig allNatural _ _ (funExt eq-l)
isIdentity
: (NT[_∘_] {A} {B} {B} (NT.identity B) f) f
× (NT[_∘_] {A} {A} {B} f (NT.identity A)) f
isIdentity = ident-l , ident-r
-- Functor categories. Objects are functors, arrows are natural transformations.
RawFun : RawCategory (c c' d d') (c c' d')
RawFun = record
{ Object = Functor 𝔻
; Arrow = NaturalTransformation
; 𝟙 = λ {F} NT.identity F
; _∘_ = λ {F G H} NT[_∘_] {F} {G} {H}
}
open RawCategory RawFun
open Univalence (λ {A} {B} {f} isIdentity {A} {B} {f})
private
module _ (F : Functor 𝔻) where module _ (F : Functor 𝔻) where
center : Σ[ G Object ] (F G) center : Σ[ G Object ] (F G)
center = F , id-to-iso F F refl center = F , id-to-iso F F refl
@ -126,7 +90,6 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
univalent[Contr] : isContr (Σ[ G Object ] (F G)) univalent[Contr] : isContr (Σ[ G Object ] (F G))
univalent[Contr] = center , isContractible univalent[Contr] = center , isContractible
private
module _ {A B : Functor 𝔻} where module _ {A B : Functor 𝔻} where
module A = Functor A module A = Functor A
module B = Functor B module B = Functor B
@ -176,69 +139,67 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
fromEq : NaturalTransformation A B fromEq : NaturalTransformation A B
fromEq = coe𝟙 , nat fromEq = coe𝟙 , nat
module _ {A B : Functor 𝔻} where module _ {A B : Functor 𝔻} where
obverse : A B A B obverse : A B A B
obverse p = res obverse p = res
where where
ob : Arrow A B ob : Arrow A B
ob = fromEq p ob = fromEq p
re : Arrow B A re : Arrow B A
re = fromEq (sym p) re = fromEq (sym p)
vr : _∘_ {A = A} {B} {A} re ob 𝟙 {A} vr : _∘_ {A = A} {B} {A} re ob 𝟙 {A}
vr = {!!} vr = {!!}
rv : _∘_ {A = B} {A} {B} ob re 𝟙 {B} rv : _∘_ {A = B} {A} {B} ob re 𝟙 {B}
rv = {!!} rv = {!!}
isInverse : IsInverseOf {A} {B} ob re isInverse : IsInverseOf {A} {B} ob re
isInverse = vr , rv isInverse = vr , rv
iso : Isomorphism {A} {B} ob iso : Isomorphism {A} {B} ob
iso = re , isInverse iso = re , isInverse
res : A B res : A B
res = ob , iso res = ob , iso
reverse : A B A B reverse : A B A B
reverse iso = {!!} reverse iso = {!!}
ve-re : (y : A B) obverse (reverse y) y ve-re : (y : A B) obverse (reverse y) y
ve-re = {!!} ve-re = {!!}
re-ve : (x : A B) reverse (obverse x) x re-ve : (x : A B) reverse (obverse x) x
re-ve = {!!} re-ve = {!!}
done : isEquiv (A B) (A B) (Univalence.id-to-iso (λ { {A} {B} isIdentity {A} {B}}) A B) done : isEquiv (A B) (A B) (Univalence.id-to-iso (λ { {A} {B} isIdentity {F = A} {B}}) A B)
done = {!gradLemma obverse reverse ve-re re-ve!} done = {!gradLemma obverse reverse ve-re re-ve!}
univalent : Univalent -- univalent : Univalent
univalent = done -- univalent = done
instance isCategory : IsCategory raw
isCategory : IsCategory RawFun IsCategory.isAssociative isCategory {A} {B} {C} {D} = isAssociative {A} {B} {C} {D}
isCategory = record IsCategory.isIdentity isCategory {A} {B} = isIdentity {A} {B}
{ isAssociative = λ {A B C D} isAssociative {A} {B} {C} {D} IsCategory.arrowsAreSets isCategory {F} {G} = naturalTransformationIsSet {F} {G}
; isIdentity = λ {A B} isIdentity {A} {B} IsCategory.univalent isCategory = {!!}
; arrowsAreSets = λ {F} {G} naturalTransformationIsSet {F} {G}
; univalent = univalent
}
Fun : Category (c c' d d') (c c' d') Fun : Category (c c' d d') (c c' d')
Category.raw Fun = RawFun Category.raw Fun = raw
Category.isCategory Fun = isCategory
module _ { ' : Level} ( : Category ') where -- module _ { ' : Level} ( : Category ') where
private -- private
open import Cat.Categories.Sets -- open import Cat.Categories.Sets
open NaturalTransformation (opposite ) (𝓢𝓮𝓽 ') -- open NaturalTransformation (opposite ) (𝓢𝓮𝓽 ')
-- Restrict the functors to Presheafs. -- -- Restrict the functors to Presheafs.
rawPresh : RawCategory ( lsuc ') ( ') -- rawPresh : RawCategory ( ⊔ lsuc ') (')
rawPresh = record -- rawPresh = record
{ Object = Presheaf -- { Object = Presheaf
; Arrow = NaturalTransformation -- ; Arrow = NaturalTransformation
; 𝟙 = λ {F} identity F -- ; 𝟙 = λ {F} → identity F
; _∘_ = λ {F G H} NT[_∘_] {F = F} {G = G} {H = H} -- ; _∘_ = λ {F G H} → NT[_∘_] {F = F} {G = G} {H = H}
} -- }
instance -- instance
isCategory : IsCategory rawPresh -- isCategory : IsCategory rawPresh
isCategory = Fun.isCategory _ _ -- isCategory = Fun.isCategory _ _
Presh : Category ( lsuc ') ( ') -- Presh : Category ( ⊔ lsuc ') (')
Category.raw Presh = rawPresh -- Category.raw Presh = rawPresh
Category.isCategory Presh = isCategory -- Category.isCategory Presh = isCategory

View file

@ -23,7 +23,7 @@ module Cat.Category.Monad where
open import Cat.Prelude open import Cat.Prelude
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor as F open import Cat.Category.Functor as F
open import Cat.Category.NaturalTransformation import Cat.Category.NaturalTransformation
import Cat.Category.Monad.Monoidal import Cat.Category.Monad.Monoidal
import Cat.Category.Monad.Kleisli import Cat.Category.Monad.Kleisli
open import Cat.Categories.Fun open import Cat.Categories.Fun
@ -33,6 +33,7 @@ module Kleisli = Cat.Category.Monad.Kleisli
-- | The monoidal- and kleisli presentation of monads are equivalent. -- | The monoidal- and kleisli presentation of monads are equivalent.
module _ {a b : Level} ( : Category a b) where module _ {a b : Level} ( : Category a b) where
open Cat.Category.NaturalTransformation
private private
module = Category module = Category
open using (Object ; Arrow ; 𝟙 ; _∘_ ; _>>>_) open using (Object ; Arrow ; 𝟙 ; _∘_ ; _>>>_)
@ -171,8 +172,6 @@ module _ {a b : Level} ( : Category a b) where
Req : M.RawMonad.R (backRaw (forth m)) R Req : M.RawMonad.R (backRaw (forth m)) R
Req = Functor≡ rawEq Req = Functor≡ rawEq
open NaturalTransformation
pureTEq : M.RawMonad.pureT (backRaw (forth m)) pureT pureTEq : M.RawMonad.pureT (backRaw (forth m)) pureT
pureTEq = funExt (λ X refl) pureTEq = funExt (λ X refl)

View file

@ -8,11 +8,11 @@ open import Cat.Prelude
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor as F open import Cat.Category.Functor as F
open import Cat.Category.NaturalTransformation
open import Cat.Categories.Fun open import Cat.Categories.Fun
-- "A monad in the Kleisli form" [voe] -- "A monad in the Kleisli form" [voe]
module Cat.Category.Monad.Kleisli {a b : Level} ( : Category a b) where module Cat.Category.Monad.Kleisli {a b : Level} ( : Category a b) where
open import Cat.Category.NaturalTransformation hiding (propIsNatural)
private private
= a b = a b
module = Category module = Category
@ -116,8 +116,6 @@ record IsMonad (raw : RawMonad) : Set where
Functor.isFunctor R = isFunctorR Functor.isFunctor R = isFunctorR
private private
open NaturalTransformation
R⁰ : EndoFunctor R⁰ : EndoFunctor
R⁰ = Functors.identity R⁰ = Functors.identity
: EndoFunctor : EndoFunctor

View file

@ -8,7 +8,6 @@ open import Cat.Prelude
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor as F open import Cat.Category.Functor as F
open import Cat.Category.NaturalTransformation
open import Cat.Categories.Fun open import Cat.Categories.Fun
module Cat.Category.Monad.Monoidal {a b : Level} ( : Category a b) where module Cat.Category.Monad.Monoidal {a b : Level} ( : Category a b) where
@ -18,7 +17,7 @@ private
= a b = a b
open Category using (Object ; Arrow ; 𝟙 ; _∘_) open Category using (Object ; Arrow ; 𝟙 ; _∘_)
open NaturalTransformation open import Cat.Category.NaturalTransformation
record RawMonad : Set where record RawMonad : Set where
field field
R : EndoFunctor R : EndoFunctor

View file

@ -9,17 +9,17 @@ open import Function
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor as F open import Cat.Category.Functor as F
open import Cat.Category.NaturalTransformation import Cat.Category.NaturalTransformation
open import Cat.Category.Monad open import Cat.Category.Monad
open import Cat.Categories.Fun open import Cat.Categories.Fun
open import Cat.Equivalence open import Cat.Equivalence
module voe {a b : Level} ( : Category a b) where module voe {a b : Level} ( : Category a b) where
open Cat.Category.NaturalTransformation
private private
= a b = a b
module = Category module = Category
open using (Object ; Arrow) open using (Object ; Arrow)
open NaturalTransformation
module M = Monoidal module M = Monoidal
module K = Kleisli module K = Kleisli

View file

@ -18,8 +18,6 @@
-- --
-- * A composition operator. -- * A composition operator.
{-# OPTIONS --allow-unsolved-metas --cubical #-} {-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Category.NaturalTransformation where
open import Cat.Prelude open import Cat.Prelude
open import Data.Nat using (_≤_ ; z≤n ; s≤s) open import Data.Nat using (_≤_ ; z≤n ; s≤s)
@ -29,77 +27,79 @@ open import Cat.Category
open import Cat.Category.Functor open import Cat.Category.Functor
open import Cat.Wishlist open import Cat.Wishlist
module NaturalTransformation {c c' d d' : Level} module Cat.Category.NaturalTransformation
{c c' d d' : Level}
( : Category c c') (𝔻 : Category d d') where ( : Category c c') (𝔻 : Category d d') where
open Category using (Object ; 𝟙) open Category using (Object ; 𝟙)
private
module = Category
module 𝔻 = Category 𝔻
module _ (F G : Functor 𝔻) where
private private
module = Category module F = Functor F
module 𝔻 = Category 𝔻 module G = Functor G
-- What do you call a non-natural tranformation?
Transformation : Set (c d')
Transformation = (C : Object ) 𝔻 [ F.omap C , G.omap C ]
module _ (F G : Functor 𝔻) where Natural : Transformation Set (c (c' d'))
private Natural θ
module F = Functor F = {A B : Object }
module G = Functor G (f : [ A , B ])
-- What do you call a non-natural tranformation? 𝔻 [ θ B F.fmap f ] 𝔻 [ G.fmap f θ A ]
Transformation : Set (c d')
Transformation = (C : Object ) 𝔻 [ F.omap C , G.omap C ]
Natural : Transformation Set (c (c' d')) NaturalTransformation : Set (c c' d')
Natural θ NaturalTransformation = Σ Transformation Natural
= {A B : Object }
(f : [ A , B ])
𝔻 [ θ B F.fmap f ] 𝔻 [ G.fmap f θ A ]
NaturalTransformation : Set (c c' d') -- Think I need propPi and that arrows are sets
NaturalTransformation = Σ Transformation Natural propIsNatural : (θ : _) isProp (Natural θ)
propIsNatural θ x y i {A} {B} f = 𝔻.arrowsAreSets _ _ (x f) (y f) i
-- Think I need propPi and that arrows are sets NaturalTransformation≡ : {α β : NaturalTransformation}
propIsNatural : (θ : _) isProp (Natural θ) (eq₁ : α .proj₁ β .proj₁)
propIsNatural θ x y i {A} {B} f = 𝔻.arrowsAreSets _ _ (x f) (y f) i α β
NaturalTransformation≡ eq = lemSig propIsNatural _ _ eq
NaturalTransformation≡ : {α β : NaturalTransformation} identityTrans : (F : Functor 𝔻) Transformation F F
(eq₁ : α .proj₁ β .proj₁) identityTrans F C = 𝟙 𝔻
α β
NaturalTransformation≡ eq = lemSig propIsNatural _ _ eq
identityTrans : (F : Functor 𝔻) Transformation F F identityNatural : (F : Functor 𝔻) Natural F F (identityTrans F)
identityTrans F C = 𝟙 𝔻 identityNatural F {A = A} {B = B} f = begin
𝔻 [ identityTrans F B F→ f ] ≡⟨⟩
𝔻 [ 𝟙 𝔻 F→ f ] ≡⟨ 𝔻.leftIdentity
F→ f ≡⟨ sym 𝔻.rightIdentity
𝔻 [ F→ f 𝟙 𝔻 ] ≡⟨⟩
𝔻 [ F→ f identityTrans F A ]
where
module F = Functor F
F→ = F.fmap
identityNatural : (F : Functor 𝔻) Natural F F (identityTrans F) identity : (F : Functor 𝔻) NaturalTransformation F F
identityNatural F {A = A} {B = B} f = begin identity F = identityTrans F , identityNatural F
𝔻 [ identityTrans F B F→ f ] ≡⟨⟩
𝔻 [ 𝟙 𝔻 F→ f ] ≡⟨ 𝔻.leftIdentity
F→ f ≡⟨ sym 𝔻.rightIdentity
𝔻 [ F→ f 𝟙 𝔻 ] ≡⟨⟩
𝔻 [ F→ f identityTrans F A ]
where
module F = Functor F
F→ = F.fmap
identity : (F : Functor 𝔻) NaturalTransformation F F module _ {F G H : Functor 𝔻} where
identity F = identityTrans F , identityNatural F private
module F = Functor F
module G = Functor G
module H = Functor H
T[_∘_] : Transformation G H Transformation F G Transformation F H
T[ θ η ] C = 𝔻 [ θ C η C ]
module _ {F G H : Functor 𝔻} where NT[_∘_] : NaturalTransformation G H NaturalTransformation F G NaturalTransformation F H
private proj₁ NT[ (θ , _) (η , _) ] = T[ θ η ]
module F = Functor F proj₂ NT[ (θ , θNat) (η , ηNat) ] {A} {B} f = begin
module G = Functor G 𝔻 [ T[ θ η ] B F.fmap f ] ≡⟨⟩
module H = Functor H 𝔻 [ 𝔻 [ θ B η B ] F.fmap f ] ≡⟨ sym 𝔻.isAssociative
T[_∘_] : Transformation G H Transformation F G Transformation F H 𝔻 [ θ B 𝔻 [ η B F.fmap f ] ] ≡⟨ cong (λ φ 𝔻 [ θ B φ ]) (ηNat f)
T[ θ η ] C = 𝔻 [ θ C η C ] 𝔻 [ θ B 𝔻 [ G.fmap f η A ] ] ≡⟨ 𝔻.isAssociative
𝔻 [ 𝔻 [ θ B G.fmap f ] η A ] ≡⟨ cong (λ φ 𝔻 [ φ η A ]) (θNat f)
NT[_∘_] : NaturalTransformation G H NaturalTransformation F G NaturalTransformation F H 𝔻 [ 𝔻 [ H.fmap f θ A ] η A ] ≡⟨ sym 𝔻.isAssociative
proj₁ NT[ (θ , _) (η , _) ] = T[ θ η ] 𝔻 [ H.fmap f 𝔻 [ θ A η A ] ] ≡⟨⟩
proj₂ NT[ (θ , θNat) (η , ηNat) ] {A} {B} f = begin 𝔻 [ H.fmap f T[ θ η ] A ]
𝔻 [ T[ θ η ] B F.fmap f ] ≡⟨⟩
𝔻 [ 𝔻 [ θ B η B ] F.fmap f ] ≡⟨ sym 𝔻.isAssociative
𝔻 [ θ B 𝔻 [ η B F.fmap f ] ] ≡⟨ cong (λ φ 𝔻 [ θ B φ ]) (ηNat f)
𝔻 [ θ B 𝔻 [ G.fmap f η A ] ] ≡⟨ 𝔻.isAssociative
𝔻 [ 𝔻 [ θ B G.fmap f ] η A ] ≡⟨ cong (λ φ 𝔻 [ φ η A ]) (θNat f)
𝔻 [ 𝔻 [ H.fmap f θ A ] η A ] ≡⟨ sym 𝔻.isAssociative
𝔻 [ H.fmap f 𝔻 [ θ A η A ] ] ≡⟨⟩
𝔻 [ H.fmap f T[ θ η ] A ]
module Properties where
module _ {F G : Functor 𝔻} where module _ {F G : Functor 𝔻} where
transformationIsSet : isSet (Transformation F G) transformationIsSet : isSet (Transformation F G)
transformationIsSet _ _ p q i j C = 𝔻.arrowsAreSets _ _ (λ l p l C) (λ l q l C) i j transformationIsSet _ _ p q i j C = 𝔻.arrowsAreSets _ _ (λ l p l C) (λ l q l C) i j
@ -118,3 +118,31 @@ module NaturalTransformation {c c' d d' : Level}
naturalTransformationIsSet : isSet (NaturalTransformation F G) naturalTransformationIsSet : isSet (NaturalTransformation F G)
naturalTransformationIsSet = sigPresSet transformationIsSet naturalIsSet naturalTransformationIsSet = sigPresSet transformationIsSet naturalIsSet
module _
{F G H I : Functor 𝔻}
{θ : NaturalTransformation F G}
{η : NaturalTransformation G H}
{ζ : NaturalTransformation H I} where
-- isAssociative : NT[ ζ ∘ NT[ η ∘ θ ] ] ≡ NT[ NT[ ζ ∘ η ] ∘ θ ]
isAssociative
: NT[_∘_] {F} {H} {I} ζ (NT[_∘_] {F} {G} {H} η θ)
NT[_∘_] {F} {G} {I} (NT[_∘_] {G} {H} {I} ζ η) θ
isAssociative
= lemSig (naturalIsProp {F = F} {I}) _ _
(funExt (λ _ 𝔻.isAssociative))
module _ {F G : Functor 𝔻} {θNT : NaturalTransformation F G} where
private
propNat = naturalIsProp {F = F} {G}
rightIdentity : (NT[_∘_] {F} {F} {G} θNT (identity F)) θNT
rightIdentity = lemSig propNat _ _ (funExt (λ _ 𝔻.rightIdentity))
leftIdentity : (NT[_∘_] {F} {G} {G} (identity G) θNT) θNT
leftIdentity = lemSig propNat _ _ (funExt (λ _ 𝔻.leftIdentity))
isIdentity
: (NT[_∘_] {F} {G} {G} (identity G) θNT) θNT
× (NT[_∘_] {F} {F} {G} θNT (identity F)) θNT
isIdentity = leftIdentity , rightIdentity

View file

@ -6,8 +6,11 @@ open import Cat.Prelude
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor open import Cat.Category.Functor
open import Cat.Category.NaturalTransformation
renaming (module Properties to F)
using ()
open import Cat.Categories.Fun open import Cat.Categories.Fun using (module Fun)
open import Cat.Categories.Sets hiding (presheaf) open import Cat.Categories.Sets hiding (presheaf)
-- There is no (small) category of categories. So we won't use _⇑_ from -- There is no (small) category of categories. So we won't use _⇑_ from
@ -47,10 +50,11 @@ module _ { : Level} { : Category } where
open RawFunctor rawYoneda hiding (fmap) open RawFunctor rawYoneda hiding (fmap)
isIdentity : IsIdentity isIdentity : IsIdentity
isIdentity {c} = lemSig (naturalIsProp {F = presheaf c} {presheaf c}) _ _ eq isIdentity {c} = lemSig prp _ _ eq
where where
eq : (λ C x [ .𝟙 x ]) identityTrans (presheaf c) eq : (λ C x [ .𝟙 x ]) identityTrans (presheaf c)
eq = funExt λ A funExt λ B .leftIdentity eq = funExt λ A funExt λ B .leftIdentity
prp = F.naturalIsProp _ _ {F = presheaf c} {presheaf c}
isDistributive : IsDistributive isDistributive : IsDistributive
isDistributive {A} {B} {C} {f = f} {g} isDistributive {A} {B} {C} {f = f} {g}