Move properties about natural transformations to that module
This commit is contained in:
parent
ef688202a2
commit
d3864dbae5
|
@ -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 , ℂ.𝟙) ≡⟨⟩
|
||||||
|
|
|
@ -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
|
|
||||||
module _ {A B : Functor ℂ 𝔻} {f : NaturalTransformation A B} where
|
|
||||||
allNatural = naturalIsProp {F = A} {B}
|
|
||||||
f' = proj₁ f
|
|
||||||
eq-r : ∀ C → (𝔻 [ f' C ∘ identityTrans A C ]) ≡ f' C
|
|
||||||
eq-r C = begin
|
|
||||||
𝔻 [ f' C ∘ identityTrans A C ] ≡⟨⟩
|
|
||||||
𝔻 [ f' C ∘ 𝔻.𝟙 ] ≡⟨ 𝔻.rightIdentity ⟩
|
|
||||||
f' C ∎
|
|
||||||
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.
|
-- Functor categories. Objects are functors, arrows are natural transformations.
|
||||||
RawFun : RawCategory (ℓc ⊔ ℓc' ⊔ ℓd ⊔ ℓd') (ℓc ⊔ ℓc' ⊔ ℓd')
|
raw : RawCategory (ℓc ⊔ ℓc' ⊔ ℓd ⊔ ℓd') (ℓc ⊔ ℓc' ⊔ ℓd')
|
||||||
RawFun = record
|
RawCategory.Object raw = Functor ℂ 𝔻
|
||||||
{ Object = Functor ℂ 𝔻
|
RawCategory.Arrow raw = NaturalTransformation
|
||||||
; Arrow = NaturalTransformation
|
RawCategory.𝟙 raw {F} = identity F
|
||||||
; 𝟙 = λ {F} → NT.identity F
|
RawCategory._∘_ raw {F} {G} {H} = NT[_∘_] {F} {G} {H}
|
||||||
; _∘_ = λ {F G H} → NT[_∘_] {F} {G} {H}
|
|
||||||
}
|
open RawCategory raw
|
||||||
|
open Univalence (λ {A} {B} {f} → isIdentity {F = A} {B} {f})
|
||||||
|
|
||||||
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
|
||||||
|
@ -204,41 +167,39 @@ module Fun {ℓc ℓc' ℓd ℓd' : Level} (ℂ : Category ℓc ℓc') (𝔻 : C
|
||||||
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
|
||||||
|
|
|
@ -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)
|
||||||
|
|
||||||
|
|
|
@ -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
|
||||||
R² : EndoFunctor ℂ
|
R² : EndoFunctor ℂ
|
||||||
|
|
|
@ -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 ℂ
|
||||||
|
|
|
@ -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 ℂ
|
||||||
|
|
||||||
|
|
|
@ -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,15 +27,16 @@ 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
|
private
|
||||||
module ℂ = Category ℂ
|
module ℂ = Category ℂ
|
||||||
module 𝔻 = Category 𝔻
|
module 𝔻 = Category 𝔻
|
||||||
|
|
||||||
module _ (F G : Functor ℂ 𝔻) where
|
module _ (F G : Functor ℂ 𝔻) where
|
||||||
private
|
private
|
||||||
module F = Functor F
|
module F = Functor F
|
||||||
module G = Functor G
|
module G = Functor G
|
||||||
|
@ -63,11 +62,11 @@ module NaturalTransformation {ℓc ℓc' ℓd ℓd' : Level}
|
||||||
→ α ≡ β
|
→ α ≡ β
|
||||||
NaturalTransformation≡ eq = lemSig propIsNatural _ _ eq
|
NaturalTransformation≡ eq = lemSig propIsNatural _ _ eq
|
||||||
|
|
||||||
identityTrans : (F : Functor ℂ 𝔻) → Transformation F F
|
identityTrans : (F : Functor ℂ 𝔻) → Transformation F F
|
||||||
identityTrans F C = 𝟙 𝔻
|
identityTrans F C = 𝟙 𝔻
|
||||||
|
|
||||||
identityNatural : (F : Functor ℂ 𝔻) → Natural F F (identityTrans F)
|
identityNatural : (F : Functor ℂ 𝔻) → Natural F F (identityTrans F)
|
||||||
identityNatural F {A = A} {B = B} f = begin
|
identityNatural F {A = A} {B = B} f = begin
|
||||||
𝔻 [ identityTrans F B ∘ F→ f ] ≡⟨⟩
|
𝔻 [ identityTrans F B ∘ F→ f ] ≡⟨⟩
|
||||||
𝔻 [ 𝟙 𝔻 ∘ F→ f ] ≡⟨ 𝔻.leftIdentity ⟩
|
𝔻 [ 𝟙 𝔻 ∘ F→ f ] ≡⟨ 𝔻.leftIdentity ⟩
|
||||||
F→ f ≡⟨ sym 𝔻.rightIdentity ⟩
|
F→ f ≡⟨ sym 𝔻.rightIdentity ⟩
|
||||||
|
@ -77,10 +76,10 @@ module NaturalTransformation {ℓc ℓc' ℓd ℓd' : Level}
|
||||||
module F = Functor F
|
module F = Functor F
|
||||||
F→ = F.fmap
|
F→ = F.fmap
|
||||||
|
|
||||||
identity : (F : Functor ℂ 𝔻) → NaturalTransformation F F
|
identity : (F : Functor ℂ 𝔻) → NaturalTransformation F F
|
||||||
identity F = identityTrans F , identityNatural F
|
identity F = identityTrans F , identityNatural F
|
||||||
|
|
||||||
module _ {F G H : Functor ℂ 𝔻} where
|
module _ {F G H : Functor ℂ 𝔻} where
|
||||||
private
|
private
|
||||||
module F = Functor F
|
module F = Functor F
|
||||||
module G = Functor G
|
module G = Functor G
|
||||||
|
@ -100,6 +99,7 @@ module NaturalTransformation {ℓc ℓc' ℓd ℓd' : Level}
|
||||||
𝔻 [ H.fmap f ∘ 𝔻 [ θ A ∘ η A ] ] ≡⟨⟩
|
𝔻 [ H.fmap f ∘ 𝔻 [ θ A ∘ η A ] ] ≡⟨⟩
|
||||||
𝔻 [ H.fmap f ∘ T[ θ ∘ η ] 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
|
||||||
|
|
|
@ -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}
|
||||||
|
|
Loading…
Reference in a new issue