Rename identity in category to ascii-name

This commit is contained in:
Frederik Hanghøj Iversen 2018-04-03 11:36:09 +02:00
parent 41b442c0d8
commit 1c6d9ad2b5
16 changed files with 221 additions and 198 deletions

View file

@ -15,10 +15,10 @@ open import Cat.Categories.Fun
-- The category of categories -- The category of categories
module _ ( ' : Level) where module _ ( ' : Level) where
RawCat : RawCategory (lsuc ( ')) ( ') RawCat : RawCategory (lsuc ( ')) ( ')
RawCategory.Object RawCat = Category ' RawCategory.Object RawCat = Category '
RawCategory.Arrow RawCat = Functor RawCategory.Arrow RawCat = Functor
RawCategory.𝟙 RawCat = Functors.identity RawCategory.identity RawCat = Functors.identity
RawCategory._∘_ RawCat = F[_∘_] RawCategory._∘_ RawCat = F[_∘_]
-- NB! `ArrowsAreSets RawCat` is *not* provable. The type of functors, -- NB! `ArrowsAreSets RawCat` is *not* provable. The type of functors,
-- however, form a groupoid! Therefore there is no (1-)category of -- however, form a groupoid! Therefore there is no (1-)category of
@ -48,8 +48,8 @@ module CatProduct { ' : Level} ( 𝔻 : Category ') where
Obj = .Object × 𝔻.Object Obj = .Object × 𝔻.Object
Arr : Obj Obj Set ' Arr : Obj Obj Set '
Arr (c , d) (c' , d') = [ c , c' ] × 𝔻 [ d , d' ] Arr (c , d) (c' , d') = [ c , c' ] × 𝔻 [ d , d' ]
𝟙 : {o : Obj} Arr o o identity : {o : Obj} Arr o o
𝟙 = .𝟙 , 𝔻.𝟙 identity = .identity , 𝔻.identity
_∘_ : _∘_ :
{a b c : Obj} {a b c : Obj}
Arr b c Arr b c
@ -58,16 +58,16 @@ module CatProduct { ' : Level} ( 𝔻 : Category ') where
_∘_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) [ bc∈C ab∈C ] , 𝔻 [ bc∈D ab∈D ]} _∘_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) [ bc∈C ab∈C ] , 𝔻 [ bc∈D ab∈D ]}
rawProduct : RawCategory ' rawProduct : RawCategory '
RawCategory.Object rawProduct = Obj RawCategory.Object rawProduct = Obj
RawCategory.Arrow rawProduct = Arr RawCategory.Arrow rawProduct = Arr
RawCategory.𝟙 rawProduct = 𝟙 RawCategory.identity rawProduct = identity
RawCategory._∘_ rawProduct = _∘_ RawCategory._∘_ rawProduct = _∘_
open RawCategory rawProduct open RawCategory rawProduct
arrowsAreSets : ArrowsAreSets arrowsAreSets : ArrowsAreSets
arrowsAreSets = setSig {sA = .arrowsAreSets} {sB = λ x 𝔻.arrowsAreSets} arrowsAreSets = setSig {sA = .arrowsAreSets} {sB = λ x 𝔻.arrowsAreSets}
isIdentity : IsIdentity 𝟙 isIdentity : IsIdentity identity
isIdentity isIdentity
= Σ≡ (fst .isIdentity) (fst 𝔻.isIdentity) = Σ≡ (fst .isIdentity) (fst 𝔻.isIdentity)
, Σ≡ (snd .isIdentity) (snd 𝔻.isIdentity) , Σ≡ (snd .isIdentity) (snd 𝔻.isIdentity)
@ -202,14 +202,14 @@ 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} (identityNT F , .𝟙 {A = snd c}) 𝔻.𝟙 ident : fmap {c} {c} (identityNT F , .identity {A = snd c}) 𝔻.identity
ident = begin ident = begin
fmap {c} {c} (Category.𝟙 (object ) {c}) ≡⟨⟩ fmap {c} {c} (Category.identity (object ) {c}) ≡⟨⟩
fmap {c} {c} (idN F , .𝟙) ≡⟨⟩ fmap {c} {c} (idN F , .identity) ≡⟨⟩
𝔻 [ identityTrans F C F.fmap .𝟙 ] ≡⟨⟩ 𝔻 [ identityTrans F C F.fmap .identity ] ≡⟨⟩
𝔻 [ 𝔻.𝟙 F.fmap .𝟙 ] ≡⟨ 𝔻.leftIdentity 𝔻 [ 𝔻.identity F.fmap .identity ] ≡⟨ 𝔻.leftIdentity
F.fmap .𝟙 ≡⟨ F.isIdentity F.fmap .identity ≡⟨ F.isIdentity
𝔻.𝟙 𝔻.identity
where where
module F = Functor F module F = Functor F
@ -278,16 +278,16 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
transpose : Functor 𝔸 object transpose : Functor 𝔸 object
eq : F[ eval (parallelProduct transpose (Functors.identity { = })) ] F eq : F[ eval (parallelProduct transpose (Functors.identity { = })) ] F
-- eq : F[ :eval: ∘ {!!} ] ≡ F -- eq : F[ :eval: ∘ {!!} ] ≡ F
-- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (𝟙 Cat {o = })) ] ≡ F -- eq : Cat [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (identity Cat {o = })) ] ≡ F
-- eq' : (Cat [ :eval: ∘ -- eq' : (Cat [ :eval: ∘
-- (record { product = product } HasProducts.|×| transpose) -- (record { product = product } HasProducts.|×| transpose)
-- (𝟙 Cat) -- (identity Cat)
-- ]) -- ])
-- ≡ F -- ≡ F
-- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758` -- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758`
-- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Cat [ -- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Cat [
-- :eval: (parallelProduct F~ (𝟙 Cat {o = }))] F) catTranspose = -- :eval: (parallelProduct F~ (identity Cat {o = }))] F) catTranspose =
-- transpose , eq -- transpose , eq
-- We don't care about filling out the holes below since they are anyways hidden -- We don't care about filling out the holes below since they are anyways hidden

View file

@ -67,7 +67,7 @@ module _ { ' : Level} (Ns : Set ) where
Raw : RawCategory -- o (lsuc lzero ⊔ o) Raw : RawCategory -- o (lsuc lzero ⊔ o)
Raw.Object Raw = FiniteDecidableSubset Raw.Object Raw = FiniteDecidableSubset
Raw.Arrow Raw = Hom Raw.Arrow Raw = Hom
Raw.𝟙 Raw {o} = inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} } Raw.identity Raw {o} = inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} }
Raw._∘_ Raw = {!!} Raw._∘_ Raw = {!!}
postulate IsCategory : IsCategory Raw postulate IsCategory : IsCategory Raw

View file

@ -11,9 +11,9 @@ module _ (a b : Level) where
Object = Σ[ hA hSet a ] (proj₁ hA hSet b) Object = Σ[ hA hSet a ] (proj₁ hA hSet b)
Arr : Object Object Set (a b) Arr : Object Object Set (a b)
Arr ((A , _) , B) ((A' , _) , B') = Σ[ f (A A') ] ({x : A} proj₁ (B x) proj₁ (B' (f x))) Arr ((A , _) , B) ((A' , _) , B') = Σ[ f (A A') ] ({x : A} proj₁ (B x) proj₁ (B' (f x)))
𝟙 : {A : Object} Arr A A identity : {A : Object} Arr A A
proj₁ 𝟙 = λ x x proj₁ identity = λ x x
proj₂ 𝟙 = λ b b proj₂ identity = λ b b
_∘_ : {a b c : Object} Arr b c Arr a b Arr a c _∘_ : {a b c : Object} Arr b c Arr a b Arr a c
(g , g') (f , f') = g Function.∘ f , g' Function.∘ f' (g , g') (f , f') = g Function.∘ f , g' Function.∘ f'
@ -21,16 +21,16 @@ module _ (a b : Level) where
RawFam = record RawFam = record
{ Object = Object { Object = Object
; Arrow = Arr ; Arrow = Arr
; 𝟙 = λ { {A} 𝟙 {A = A}} ; identity = λ { {A} identity {A = A}}
; _∘_ = λ {a b c} _∘_ {a} {b} {c} ; _∘_ = λ {a b c} _∘_ {a} {b} {c}
} }
open RawCategory RawFam hiding (Object ; 𝟙) open RawCategory RawFam hiding (Object ; identity)
isAssociative : IsAssociative isAssociative : IsAssociative
isAssociative = Σ≡ refl refl isAssociative = Σ≡ refl refl
isIdentity : IsIdentity λ { {A} 𝟙 {A} } isIdentity : IsIdentity λ { {A} identity {A} }
isIdentity = (Σ≡ refl refl) , Σ≡ refl refl isIdentity = (Σ≡ refl refl) , Σ≡ refl refl
open import Cubical.NType.Properties open import Cubical.NType.Properties

View file

@ -27,10 +27,10 @@ module _ {a b : Level} ( : Category a b) where
module = Category module = Category
RawFree : RawCategory a (a b) RawFree : RawCategory a (a b)
RawCategory.Object RawFree = .Object RawCategory.Object RawFree = .Object
RawCategory.Arrow RawFree = Path .Arrow RawCategory.Arrow RawFree = Path .Arrow
RawCategory.𝟙 RawFree = empty RawCategory.identity RawFree = empty
RawCategory._∘_ RawFree = concatenate RawCategory._∘_ RawFree = concatenate
open RawCategory RawFree open RawCategory RawFree
@ -52,7 +52,7 @@ module _ {a b : Level} ( : Category a b) where
ident-l : {A} {B} {p : Path .Arrow A B} concatenate empty p p ident-l : {A} {B} {p : Path .Arrow A B} concatenate empty p p
ident-l = refl ident-l = refl
isIdentity : IsIdentity 𝟙 isIdentity : IsIdentity identity
isIdentity = ident-l , ident-r isIdentity = ident-l , ident-r
open Univalence isIdentity open Univalence isIdentity

View file

@ -1,28 +1,28 @@
{-# OPTIONS --allow-unsolved-metas --cubical #-} {-# OPTIONS --allow-unsolved-metas --cubical --caching #-}
module Cat.Categories.Fun where module Cat.Categories.Fun where
open import Cat.Prelude open import Cat.Prelude
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor open import Cat.Category.Functor
import Cat.Category.NaturalTransformation
as 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
import Cat.Category.NaturalTransformation 𝔻 open NaturalTransformation 𝔻 public hiding (module Properties)
as NaturalTransformation open NaturalTransformation.Properties 𝔻
open NaturalTransformation public hiding (module Properties)
open NaturalTransformation.Properties
private private
module = Category module = Category
module 𝔻 = Category 𝔻 module 𝔻 = Category 𝔻
-- Functor categories. Objects are functors, arrows are natural transformations. -- Functor categories. Objects are functors, arrows are natural transformations.
raw : RawCategory (c c' d d') (c c' d') raw : RawCategory (c c' d d') (c c' d')
RawCategory.Object raw = Functor 𝔻 RawCategory.Object raw = Functor 𝔻
RawCategory.Arrow raw = NaturalTransformation RawCategory.Arrow raw = NaturalTransformation
RawCategory.𝟙 raw {F} = identity F RawCategory.identity raw {F} = identity F
RawCategory._∘_ raw {F} {G} {H} = NT[_∘_] {F} {G} {H} RawCategory._∘_ raw {F} {G} {H} = NT[_∘_] {F} {G} {H}
open RawCategory raw open RawCategory raw hiding (identity)
open Univalence (λ {A} {B} {f} isIdentity {F = A} {B} {f}) open Univalence (λ {A} {B} {f} isIdentity {F = A} {B} {f})
module _ (F : Functor 𝔻) where module _ (F : Functor 𝔻) where
@ -32,7 +32,7 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
open Σ center renaming (proj₂ to isoF) open Σ center renaming (proj₂ to isoF)
module _ (cG : Σ[ G Object ] (F G)) where module _ (cG : Σ[ G Object ] (F G)) where
open Σ cG renaming (proj₁ to G ; proj₂ to isoG) open Σ cG renaming (proj₁ to G ; proj₂ to isoG)
module G = Functor G module G = Functor G
open Σ isoG renaming (proj₁ to θNT ; proj₂ to invθNT) open Σ isoG renaming (proj₁ to θNT ; proj₂ to invθNT)
open Σ invθNT renaming (proj₁ to ηNT ; proj₂ to areInv) open Σ invθNT renaming (proj₁ to ηNT ; proj₂ to areInv)
@ -46,10 +46,10 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
-- g = T[ η ∘ θ ] {!!} -- g = T[ η ∘ θ ] {!!}
ntF : NaturalTransformation F F ntF : NaturalTransformation F F
ntF = 𝟙 {A = F} ntF = identity F
ntG : NaturalTransformation G G ntG : NaturalTransformation G G
ntG = 𝟙 {A = G} ntG = identity G
idFunctor = Functors.identity idFunctor = Functors.identity
@ -103,14 +103,14 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
-- The transformation will be the identity on 𝔻. Such an arrow has the -- The transformation will be the identity on 𝔻. Such an arrow has the
-- type `A.omap A → A.omap A`. Which we can coerce to have the type -- type `A.omap A → A.omap A`. Which we can coerce to have the type
-- `A.omap → B.omap` since `A` and `B` are equal. -- `A.omap → B.omap` since `A` and `B` are equal.
coe𝟙 : Transformation A B coeidentity : Transformation A B
coe𝟙 X = coe coerceAB 𝔻.𝟙 coeidentity X = coe coerceAB 𝔻.identity
module _ {a b : .Object} (f : [ a , b ]) where module _ {a b : .Object} (f : [ a , b ]) where
nat' : 𝔻 [ coe𝟙 b A.fmap f ] 𝔻 [ B.fmap f coe𝟙 a ] nat' : 𝔻 [ coeidentity b A.fmap f ] 𝔻 [ B.fmap f coeidentity a ]
nat' = begin nat' = begin
(𝔻 [ coe𝟙 b A.fmap f ]) ≡⟨ {!!} (𝔻 [ coeidentity b A.fmap f ]) ≡⟨ {!!}
(𝔻 [ B.fmap f coe𝟙 a ]) (𝔻 [ B.fmap f coeidentity a ])
transs : (i : I) Transformation A (p i) transs : (i : I) Transformation A (p i)
transs = {!!} transs = {!!}
@ -118,26 +118,26 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
natt : (i : I) Natural A (p i) {!!} natt : (i : I) Natural A (p i) {!!}
natt = {!!} natt = {!!}
t : Natural A B coe𝟙 t : Natural A B coeidentity
t = coe c (identityNatural A) t = coe c (identityNatural A)
where where
c : Natural A A (identityTrans A) Natural A B coe𝟙 c : Natural A A (identityTrans A) Natural A B coeidentity
c = begin c = begin
Natural A A (identityTrans A) ≡⟨ (λ x {!natt ?!}) Natural A A (identityTrans A) ≡⟨ (λ x {!natt ?!})
Natural A B coe𝟙 Natural A B coeidentity
-- cong (λ φ → {!Natural A A (identityTrans A)!}) {!!} -- cong (λ φ → {!Natural A A (identityTrans A)!}) {!!}
k : Natural A A (identityTrans A) Natural A B coe𝟙 k : Natural A A (identityTrans A) Natural A B coeidentity
k n {a} {b} f = res k n {a} {b} f = res
where where
res : (𝔻 [ coe𝟙 b A.fmap f ]) (𝔻 [ B.fmap f coe𝟙 a ]) res : (𝔻 [ coeidentity b A.fmap f ]) (𝔻 [ B.fmap f coeidentity a ])
res = {!!} res = {!!}
nat : Natural A B coe𝟙 nat : Natural A B coeidentity
nat = nat' nat = nat'
fromEq : NaturalTransformation A B fromEq : NaturalTransformation A B
fromEq = coe𝟙 , nat fromEq = coeidentity , nat
module _ {A B : Functor 𝔻} where module _ {A B : Functor 𝔻} where
obverse : A B A B obverse : A B A B
@ -147,9 +147,9 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
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 identity A
vr = {!!} vr = {!!}
rv : _∘_ {A = B} {A} {B} ob re 𝟙 {B} rv : _∘_ {A = B} {A} {B} ob re identity B
rv = {!!} rv = {!!}
isInverse : IsInverseOf {A} {B} ob re isInverse : IsInverseOf {A} {B} ob re
isInverse = vr , rv isInverse = vr , rv
@ -183,23 +183,42 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
Category.raw Fun = raw Category.raw Fun = raw
Category.isCategory Fun = isCategory 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 ) (𝓢𝓮𝓽 ')
module K = Fun (opposite ) (𝓢𝓮𝓽 ')
module F = Category K.Fun
-- -- Restrict the functors to Presheafs. -- Restrict the functors to Presheafs.
-- rawPresh : RawCategory ( ⊔ lsuc ') (') raw : RawCategory ( lsuc ') ( ')
-- rawPresh = record raw = record
-- { Object = Presheaf { Object = Presheaf
-- ; Arrow = NaturalTransformation ; Arrow = NaturalTransformation
-- ; 𝟙 = λ {F} → identity F ; identity = λ {F} identity F
-- ; _∘_ = λ {F G H} → NT[_∘_] {F = F} {G = G} {H = H} ; _∘_ = λ {F G H} NT[_∘_] {F = F} {G = G} {H = H}
-- } }
-- instance
-- isCategory : IsCategory rawPresh
-- isCategory = Fun.isCategory _ _
-- Presh : Category ( ⊔ lsuc ') (') isCategory : IsCategory raw
-- Category.raw Presh = rawPresh isCategory = record
-- Category.isCategory Presh = isCategory { isAssociative =
λ{ {A} {B} {C} {D} {f} {g} {h}
F.isAssociative {A} {B} {C} {D} {f} {g} {h}
}
; isIdentity =
λ{ {A} {B} {f}
F.isIdentity {A} {B} {f}
}
; arrowsAreSets =
λ{ {A} {B}
F.arrowsAreSets {A} {B}
}
; univalent =
λ{ {A} {B}
F.univalent {A} {B}
}
}
Presh : Category ( lsuc ') ( ')
Category.raw Presh = raw
Category.isCategory Presh = isCategory

View file

@ -153,7 +153,7 @@ RawRel : RawCategory (lsuc lzero) (lsuc lzero)
RawRel = record RawRel = record
{ Object = Set { Object = Set
; Arrow = λ S R Subset (S × R) ; Arrow = λ S R Subset (S × R)
; 𝟙 = λ {S} Diag S ; identity = λ {S} Diag S
; _∘_ = λ {A B C} S R λ {( a , c ) Σ[ b B ] ( (a , b) R × (b , c) S )} ; _∘_ = λ {A B C} S R λ {( a , c ) Σ[ b B ] ( (a , b) R × (b , c) S )}
} }

View file

@ -60,10 +60,10 @@ module _ { : Level} {A B : Set } {a : A} where
module _ ( : Level) where module _ ( : Level) where
private private
SetsRaw : RawCategory (lsuc ) SetsRaw : RawCategory (lsuc )
RawCategory.Object SetsRaw = hSet RawCategory.Object SetsRaw = hSet
RawCategory.Arrow SetsRaw (T , _) (U , _) = T U RawCategory.Arrow SetsRaw (T , _) (U , _) = T U
RawCategory.𝟙 SetsRaw = Function.id RawCategory.identity SetsRaw = Function.id
RawCategory._∘_ SetsRaw = Function._∘_ RawCategory._∘_ SetsRaw = Function._∘_
open RawCategory SetsRaw hiding (_∘_) open RawCategory SetsRaw hiding (_∘_)

View file

@ -12,7 +12,7 @@
-- --
-- Data -- Data
-- ---- -- ----
-- 𝟙; the identity arrow -- identity; the identity arrow
-- _∘_; function composition -- _∘_; function composition
-- --
-- Laws -- Laws
@ -48,10 +48,10 @@ import Function
record RawCategory (a b : Level) : Set (lsuc (a b)) where record RawCategory (a b : Level) : Set (lsuc (a b)) where
no-eta-equality no-eta-equality
field field
Object : Set a Object : Set a
Arrow : Object Object Set b Arrow : Object Object Set b
𝟙 : {A : Object} Arrow A A identity : {A : Object} Arrow A A
_∘_ : {A B C : Object} Arrow B C Arrow A B Arrow A C _∘_ : {A B C : Object} Arrow B C Arrow A B Arrow A C
infixl 10 _∘_ _>>>_ infixl 10 _∘_ _>>>_
@ -82,7 +82,7 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
ArrowsAreSets = {A B : Object} isSet (Arrow A B) ArrowsAreSets = {A B : Object} isSet (Arrow A B)
IsInverseOf : {A B} (Arrow A B) (Arrow B A) Set b IsInverseOf : {A B} (Arrow A B) (Arrow B A) Set b
IsInverseOf = λ f g g f 𝟙 × f g 𝟙 IsInverseOf = λ f g g f identity × f g identity
Isomorphism : {A B} (f : Arrow A B) Set b Isomorphism : {A B} (f : Arrow A B) Set b
Isomorphism {A} {B} f = Σ[ g Arrow B A ] IsInverseOf f g Isomorphism {A} {B} f = Σ[ g Arrow B A ] IsInverseOf f g
@ -110,10 +110,10 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
Terminal = Σ Object IsTerminal Terminal = Σ Object IsTerminal
-- | Univalence is indexed by a raw category as well as an identity proof. -- | Univalence is indexed by a raw category as well as an identity proof.
module Univalence (isIdentity : IsIdentity 𝟙) where module Univalence (isIdentity : IsIdentity identity) where
-- | The identity isomorphism -- | The identity isomorphism
idIso : (A : Object) A A idIso : (A : Object) A A
idIso A = 𝟙 , 𝟙 , isIdentity idIso A = identity , identity , isIdentity
-- | Extract an isomorphism from an equality -- | Extract an isomorphism from an equality
-- --
@ -150,16 +150,16 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
open RawCategory public open RawCategory public
field field
isAssociative : IsAssociative isAssociative : IsAssociative
isIdentity : IsIdentity 𝟙 isIdentity : IsIdentity identity
arrowsAreSets : ArrowsAreSets arrowsAreSets : ArrowsAreSets
open Univalence isIdentity public open Univalence isIdentity public
field field
univalent : Univalent univalent : Univalent
leftIdentity : {A B : Object} {f : Arrow A B} 𝟙 f f leftIdentity : {A B : Object} {f : Arrow A B} identity f f
leftIdentity {A} {B} {f} = fst (isIdentity {A = A} {B} {f}) leftIdentity {A} {B} {f} = fst (isIdentity {A = A} {B} {f})
rightIdentity : {A B : Object} {f : Arrow A B} f 𝟙 f rightIdentity : {A B : Object} {f : Arrow A B} f identity f
rightIdentity {A} {B} {f} = snd (isIdentity {A = A} {B} {f}) rightIdentity {A} {B} {f} = snd (isIdentity {A = A} {B} {f})
------------ ------------
@ -171,24 +171,24 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
iso→epi : Isomorphism f Epimorphism {X = X} f iso→epi : Isomorphism f Epimorphism {X = X} f
iso→epi (f- , left-inv , right-inv) g₀ g₁ eq = begin iso→epi (f- , left-inv , right-inv) g₀ g₁ eq = begin
g₀ ≡⟨ sym rightIdentity g₀ ≡⟨ sym rightIdentity
g₀ 𝟙 ≡⟨ cong (_∘_ g₀) (sym right-inv) g₀ identity ≡⟨ cong (_∘_ g₀) (sym right-inv)
g₀ (f f-) ≡⟨ isAssociative g₀ (f f-) ≡⟨ isAssociative
(g₀ f) f- ≡⟨ cong (λ φ φ f-) eq (g₀ f) f- ≡⟨ cong (λ φ φ f-) eq
(g₁ f) f- ≡⟨ sym isAssociative (g₁ f) f- ≡⟨ sym isAssociative
g₁ (f f-) ≡⟨ cong (_∘_ g₁) right-inv g₁ (f f-) ≡⟨ cong (_∘_ g₁) right-inv
g₁ 𝟙 ≡⟨ rightIdentity g₁ identity ≡⟨ rightIdentity
g₁ g₁
iso→mono : Isomorphism f Monomorphism {X = X} f iso→mono : Isomorphism f Monomorphism {X = X} f
iso→mono (f- , left-inv , right-inv) g₀ g₁ eq = iso→mono (f- , left-inv , right-inv) g₀ g₁ eq =
begin begin
g₀ ≡⟨ sym leftIdentity g₀ ≡⟨ sym leftIdentity
𝟙 g₀ ≡⟨ cong (λ φ φ g₀) (sym left-inv) identity g₀ ≡⟨ cong (λ φ φ g₀) (sym left-inv)
(f- f) g₀ ≡⟨ sym isAssociative (f- f) g₀ ≡⟨ sym isAssociative
f- (f g₀) ≡⟨ cong (_∘_ f-) eq f- (f g₀) ≡⟨ cong (_∘_ f-) eq
f- (f g₁) ≡⟨ isAssociative f- (f g₁) ≡⟨ isAssociative
(f- f) g₁ ≡⟨ cong (λ φ φ g₁) left-inv (f- f) g₁ ≡⟨ cong (λ φ φ g₁) left-inv
𝟙 g₁ ≡⟨ leftIdentity identity g₁ ≡⟨ leftIdentity
g₁ g₁
iso→epi×mono : Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f iso→epi×mono : Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f
@ -228,12 +228,12 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
where where
geq : g g' geq : g g'
geq = begin geq = begin
g ≡⟨ sym rightIdentity g ≡⟨ sym rightIdentity
g 𝟙 ≡⟨ cong (λ φ g φ) (sym ε') g identity ≡⟨ cong (λ φ g φ) (sym ε')
g (f g') ≡⟨ isAssociative g (f g') ≡⟨ isAssociative
(g f) g' ≡⟨ cong (λ φ φ g') η (g f) g' ≡⟨ cong (λ φ φ g') η
𝟙 g' ≡⟨ leftIdentity identity g' ≡⟨ leftIdentity
g' g'
propUnivalent : isProp Univalent propUnivalent : isProp Univalent
propUnivalent a b i = propPi (λ iso propIsContr) a b i propUnivalent a b i = propPi (λ iso propIsContr) a b i
@ -274,9 +274,9 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
Xprop f g = trans (sym (snd Xit f)) (snd Xit g) Xprop f g = trans (sym (snd Xit f)) (snd Xit g)
Yprop : isProp (Arrow Y Y) Yprop : isProp (Arrow Y Y)
Yprop f g = trans (sym (snd Yit f)) (snd Yit g) Yprop f g = trans (sym (snd Yit f)) (snd Yit g)
left : Y→X X→Y 𝟙 left : Y→X X→Y identity
left = Xprop _ _ left = Xprop _ _
right : X→Y Y→X 𝟙 right : X→Y Y→X identity
right = Yprop _ _ right = Yprop _ _
iso : X Y iso : X Y
iso = X→Y , Y→X , left , right iso = X→Y , Y→X , left , right
@ -321,9 +321,9 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
Xprop f g = trans (sym (snd Xii f)) (snd Xii g) Xprop f g = trans (sym (snd Xii f)) (snd Xii g)
Yprop : isProp (Arrow Y Y) Yprop : isProp (Arrow Y Y)
Yprop f g = trans (sym (snd Yii f)) (snd Yii g) Yprop f g = trans (sym (snd Yii f)) (snd Yii g)
left : Y→X X→Y 𝟙 left : Y→X X→Y identity
left = Yprop _ _ left = Yprop _ _
right : X→Y Y→X 𝟙 right : X→Y Y→X identity
right = Xprop _ _ right = Xprop _ _
iso : X Y iso : X Y
iso = Y→X , X→Y , right , left iso = Y→X , X→Y , right , left
@ -351,18 +351,18 @@ module _ {a b : Level} ( : RawCategory a b) where
-- adverse effects this may have. -- adverse effects this may have.
module Prop = X.Propositionality module Prop = X.Propositionality
isIdentity : (λ _ IsIdentity 𝟙) [ X.isIdentity Y.isIdentity ] isIdentity : (λ _ IsIdentity identity) [ X.isIdentity Y.isIdentity ]
isIdentity = Prop.propIsIdentity X.isIdentity Y.isIdentity isIdentity = Prop.propIsIdentity X.isIdentity Y.isIdentity
U : {a : IsIdentity 𝟙} U : {a : IsIdentity identity}
(λ _ IsIdentity 𝟙) [ X.isIdentity a ] (λ _ IsIdentity identity) [ X.isIdentity a ]
(b : Univalent a) (b : Univalent a)
Set _ Set _
U eqwal univ = U eqwal univ =
(λ i Univalent (eqwal i)) (λ i Univalent (eqwal i))
[ X.univalent univ ] [ X.univalent univ ]
P : (y : IsIdentity 𝟙) P : (y : IsIdentity identity)
(λ _ IsIdentity 𝟙) [ X.isIdentity y ] Set _ (λ _ IsIdentity identity) [ X.isIdentity y ] Set _
P y eq = (univ : Univalent y) U eq univ P y eq = (univ : Univalent y) U eq univ
p : (b' : Univalent X.isIdentity) p : (b' : Univalent X.isIdentity)
(λ _ Univalent X.isIdentity) [ X.univalent b' ] (λ _ Univalent X.isIdentity) [ X.univalent b' ]
@ -426,14 +426,14 @@ module Opposite {a b : Level} where
private private
module = Category module = Category
opRaw : RawCategory a b opRaw : RawCategory a b
RawCategory.Object opRaw = .Object RawCategory.Object opRaw = .Object
RawCategory.Arrow opRaw = Function.flip .Arrow RawCategory.Arrow opRaw = Function.flip .Arrow
RawCategory.𝟙 opRaw = .𝟙 RawCategory.identity opRaw = .identity
RawCategory._∘_ opRaw = Function.flip ._∘_ RawCategory._∘_ opRaw = Function.flip ._∘_
open RawCategory opRaw open RawCategory opRaw
isIdentity : IsIdentity 𝟙 isIdentity : IsIdentity identity
isIdentity = swap .isIdentity isIdentity = swap .isIdentity
open Univalence isIdentity open Univalence isIdentity
@ -530,7 +530,7 @@ module Opposite {a b : Level} where
rawInv : Category.raw (opposite (opposite )) raw rawInv : Category.raw (opposite (opposite )) raw
RawCategory.Object (rawInv _) = Object RawCategory.Object (rawInv _) = Object
RawCategory.Arrow (rawInv _) = Arrow RawCategory.Arrow (rawInv _) = Arrow
RawCategory.𝟙 (rawInv _) = 𝟙 RawCategory.identity (rawInv _) = identity
RawCategory._∘_ (rawInv _) = _∘_ RawCategory._∘_ (rawInv _) = _∘_
oppositeIsInvolution : opposite (opposite ) oppositeIsInvolution : opposite (opposite )

View file

@ -16,11 +16,11 @@ module _ { '} ( : Category ') {{hasProducts : HasProducts }}
field field
uniq uniq
: (A : Object) (f : [ A × B , C ]) : (A : Object) (f : [ A × B , C ])
∃![ f~ ] ( [ eval f~ |×| Category.𝟙 ] f) ∃![ f~ ] ( [ eval f~ |×| Category.identity ] f)
IsExponential : (Cᴮ : Object) [ Cᴮ × B , C ] Set ( ') IsExponential : (Cᴮ : Object) [ Cᴮ × B , C ] Set ( ')
IsExponential Cᴮ eval = (A : Object) (f : [ A × B , C ]) IsExponential Cᴮ eval = (A : Object) (f : [ A × B , C ])
∃![ f~ ] ( [ eval f~ |×| Category.𝟙 ] f) ∃![ f~ ] ( [ eval f~ |×| Category.identity ] f)
record Exponential : Set ( ') where record Exponential : Set ( ') where
field field

View file

@ -30,7 +30,7 @@ module _ {c c' d d'}
fmap : {A B} [ A , B ] 𝔻 [ omap A , omap B ] fmap : {A B} [ A , B ] 𝔻 [ omap A , omap B ]
IsIdentity : Set _ IsIdentity : Set _
IsIdentity = {A : .Object} fmap (.𝟙 {A}) 𝔻.𝟙 {omap A} IsIdentity = {A : .Object} fmap (.identity {A}) 𝔻.identity {omap A}
IsDistributive : Set _ IsDistributive : Set _
IsDistributive = {A B C : .Object} {f : [ A , B ]} {g : [ B , C ]} IsDistributive = {A B C : .Object} {f : [ A , B ]} {g : [ B , C ]}
@ -150,10 +150,10 @@ module _ {0 1 2 3 4 5 : Level}
isFunctor : IsFunctor A C raw isFunctor : IsFunctor A C raw
isFunctor = record isFunctor = record
{ isIdentity = begin { isIdentity = begin
(F.fmap G.fmap) A.𝟙 ≡⟨ refl (F.fmap G.fmap) A.identity ≡⟨ refl
F.fmap (G.fmap A.𝟙) ≡⟨ cong F.fmap (G.isIdentity) F.fmap (G.fmap A.identity) ≡⟨ cong F.fmap (G.isIdentity)
F.fmap B.𝟙 ≡⟨ F.isIdentity F.fmap B.identity ≡⟨ F.isIdentity
C.𝟙 C.identity
; isDistributive = dist ; isDistributive = dist
} }

View file

@ -33,10 +33,10 @@ 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 open Cat.Category.NaturalTransformation using (NaturalTransformation ; propIsNatural)
private private
module = Category module = Category
open using (Object ; Arrow ; 𝟙 ; _∘_ ; _>>>_) open using (Object ; Arrow ; identity ; _∘_ ; _>>>_)
module M = Monoidal module M = Monoidal
module K = Kleisli module K = Kleisli
@ -84,11 +84,11 @@ module _ {a b : Level} ( : Category a b) where
inv-l = begin inv-l = begin
joinT X pureT (R.omap X) ≡⟨⟩ joinT X pureT (R.omap X) ≡⟨⟩
join pure ≡⟨ proj₁ isInverse join pure ≡⟨ proj₁ isInverse
𝟙 identity
inv-r = begin inv-r = begin
joinT X R.fmap (pureT X) ≡⟨⟩ joinT X R.fmap (pureT X) ≡⟨⟩
join fmap pure ≡⟨ proj₂ isInverse join fmap pure ≡⟨ proj₂ isInverse
𝟙 identity
back : K.Monad M.Monad back : K.Monad M.Monad
Monoidal.Monad.raw (back m) = backRaw m Monoidal.Monad.raw (back m) = backRaw m
@ -103,21 +103,21 @@ module _ {a b : Level} ( : Category a b) where
bindEq {X} {Y} = begin bindEq {X} {Y} = begin
K.RawMonad.bind (forthRaw (backRaw m)) ≡⟨⟩ K.RawMonad.bind (forthRaw (backRaw m)) ≡⟨⟩
(λ f join fmap f) ≡⟨⟩ (λ f join fmap f) ≡⟨⟩
(λ f bind (f >>> pure) >>> bind 𝟙) ≡⟨ funExt lem (λ f bind (f >>> pure) >>> bind identity) ≡⟨ funExt lem
(λ f bind f) ≡⟨⟩ (λ f bind f) ≡⟨⟩
bind bind
where where
lem : (f : Arrow X (omap Y)) lem : (f : Arrow X (omap Y))
bind (f >>> pure) >>> bind 𝟙 bind (f >>> pure) >>> bind identity
bind f bind f
lem f = begin lem f = begin
bind (f >>> pure) >>> bind 𝟙 bind (f >>> pure) >>> bind identity
≡⟨ isDistributive _ _ ≡⟨ isDistributive _ _
bind ((f >>> pure) >>> bind 𝟙) bind ((f >>> pure) >>> bind identity)
≡⟨ cong bind .isAssociative ≡⟨ cong bind .isAssociative
bind (f >>> (pure >>> bind 𝟙)) bind (f >>> (pure >>> bind identity))
≡⟨ cong (λ φ bind (f >>> φ)) (isNatural _) ≡⟨ cong (λ φ bind (f >>> φ)) (isNatural _)
bind (f >>> 𝟙) bind (f >>> identity)
≡⟨ cong bind .leftIdentity ≡⟨ cong bind .leftIdentity
bind f bind f
@ -146,10 +146,10 @@ module _ {a b : Level} ( : Category a b) where
joinEq : {X} KM.join joinT X joinEq : {X} KM.join joinT X
joinEq {X} = begin joinEq {X} = begin
KM.join ≡⟨⟩ KM.join ≡⟨⟩
KM.bind 𝟙 ≡⟨⟩ KM.bind identity ≡⟨⟩
bind 𝟙 ≡⟨⟩ bind identity ≡⟨⟩
joinT X Rfmap 𝟙 ≡⟨ cong (λ φ _ φ) R.isIdentity joinT X Rfmap identity ≡⟨ cong (λ φ _ φ) R.isIdentity
joinT X 𝟙 ≡⟨ .rightIdentity joinT X identity ≡⟨ .rightIdentity
joinT X joinT X
fmapEq : {A B} KM.fmap {A} {B} Rfmap fmapEq : {A B} KM.fmap {A} {B} Rfmap
@ -161,7 +161,7 @@ module _ {a b : Level} ( : Category a b) where
Rfmap (f >>> pureT B) >>> joinT B ≡⟨ cong (λ φ φ >>> joinT B) R.isDistributive Rfmap (f >>> pureT B) >>> joinT B ≡⟨ cong (λ φ φ >>> joinT B) R.isDistributive
Rfmap f >>> Rfmap (pureT B) >>> joinT B ≡⟨ .isAssociative Rfmap f >>> Rfmap (pureT B) >>> joinT B ≡⟨ .isAssociative
joinT B Rfmap (pureT B) Rfmap f ≡⟨ cong (λ φ φ Rfmap f) (proj₂ isInverse) joinT B Rfmap (pureT B) Rfmap f ≡⟨ cong (λ φ φ Rfmap f) (proj₂ isInverse)
𝟙 Rfmap f ≡⟨ .leftIdentity identity Rfmap f ≡⟨ .leftIdentity
Rfmap f Rfmap f
) )
@ -183,8 +183,8 @@ module _ {a b : Level} ( : Category a b) where
joinTEq = funExt (λ X begin joinTEq = funExt (λ X begin
M.RawMonad.joinT (backRaw (forth m)) X ≡⟨⟩ M.RawMonad.joinT (backRaw (forth m)) X ≡⟨⟩
KM.join ≡⟨⟩ KM.join ≡⟨⟩
joinT X Rfmap 𝟙 ≡⟨ cong (λ φ joinT X φ) R.isIdentity joinT X Rfmap identity ≡⟨ cong (λ φ joinT X φ) R.isIdentity
joinT X 𝟙 ≡⟨ .rightIdentity joinT X identity ≡⟨ .rightIdentity
joinT X ) joinT X )
joinNTEq : (λ i NaturalTransformation F[ Req i Req i ] (Req i)) joinNTEq : (λ i NaturalTransformation F[ Req i Req i ] (Req i))

View file

@ -12,11 +12,13 @@ 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) open import Cat.Category.NaturalTransformation
using (NaturalTransformation ; Transformation ; Natural)
private private
= a b = a b
module = Category module = Category
open using (Arrow ; 𝟙 ; Object ; _∘_ ; _>>>_) open using (Arrow ; identity ; Object ; _∘_ ; _>>>_)
-- | Data for a monad. -- | Data for a monad.
-- --
@ -40,7 +42,7 @@ record RawMonad : Set where
-- | Flattening nested monads. -- | Flattening nested monads.
join : {A : Object} [ omap (omap A) , omap A ] join : {A : Object} [ omap (omap A) , omap A ]
join = bind 𝟙 join = bind identity
------------------ ------------------
-- * Monad laws -- -- * Monad laws --
@ -49,7 +51,7 @@ record RawMonad : Set where
-- There may be better names than what I've chosen here. -- There may be better names than what I've chosen here.
IsIdentity = {X : Object} IsIdentity = {X : Object}
bind pure 𝟙 {omap X} bind pure identity {omap X}
IsNatural = {X Y : Object} (f : [ X , omap Y ]) IsNatural = {X Y : Object} (f : [ X , omap Y ])
pure >>> (bind f) f pure >>> (bind f) f
IsDistributive = {X Y Z : Object} (g : [ Y , omap Z ]) (f : [ X , omap Y ]) IsDistributive = {X Y Z : Object} (g : [ Y , omap Z ]) (f : [ X , omap Y ])
@ -67,7 +69,7 @@ record RawMonad : Set where
IsNaturalForeign = {X : Object} join {X} fmap join join join IsNaturalForeign = {X : Object} join {X} fmap join join join
IsInverse : Set _ IsInverse : Set _
IsInverse = {X : Object} join {X} pure 𝟙 × join {X} fmap pure 𝟙 IsInverse = {X : Object} join {X} pure identity × join {X} fmap pure identity
record IsMonad (raw : RawMonad) : Set where record IsMonad (raw : RawMonad) : Set where
open RawMonad raw public open RawMonad raw public
@ -100,9 +102,9 @@ record IsMonad (raw : RawMonad) : Set where
isFunctorR : IsFunctor rawR isFunctorR : IsFunctor rawR
IsFunctor.isIdentity isFunctorR = begin IsFunctor.isIdentity isFunctorR = begin
bind (pure 𝟙) ≡⟨ cong bind (.rightIdentity) bind (pure identity) ≡⟨ cong bind (.rightIdentity)
bind pure ≡⟨ isIdentity bind pure ≡⟨ isIdentity
𝟙 identity
IsFunctor.isDistributive isFunctorR {f = f} {g} = begin IsFunctor.isDistributive isFunctorR {f = f} {g} = begin
bind (pure (g f)) ≡⟨⟩ bind (pure (g f)) ≡⟨⟩
@ -137,29 +139,29 @@ record IsMonad (raw : RawMonad) : Set where
joinN : Natural R joinT joinN : Natural R joinT
joinN f = begin joinN f = begin
join R².fmap f ≡⟨⟩ join R².fmap f ≡⟨⟩
bind 𝟙 R².fmap f ≡⟨⟩ bind identity R².fmap f ≡⟨⟩
R².fmap f >>> bind 𝟙 ≡⟨⟩ R².fmap f >>> bind identity ≡⟨⟩
fmap (fmap f) >>> bind 𝟙 ≡⟨⟩ fmap (fmap f) >>> bind identity ≡⟨⟩
fmap (bind (f >>> pure)) >>> bind 𝟙 ≡⟨⟩ fmap (bind (f >>> pure)) >>> bind identity ≡⟨⟩
bind (bind (f >>> pure) >>> pure) >>> bind 𝟙 bind (bind (f >>> pure) >>> pure) >>> bind identity
≡⟨ isDistributive _ _ ≡⟨ isDistributive _ _
bind ((bind (f >>> pure) >>> pure) >=> 𝟙) bind ((bind (f >>> pure) >>> pure) >=> identity)
≡⟨⟩ ≡⟨⟩
bind ((bind (f >>> pure) >>> pure) >>> bind 𝟙) bind ((bind (f >>> pure) >>> pure) >>> bind identity)
≡⟨ cong bind .isAssociative ≡⟨ cong bind .isAssociative
bind (bind (f >>> pure) >>> (pure >>> bind 𝟙)) bind (bind (f >>> pure) >>> (pure >>> bind identity))
≡⟨ cong (λ φ bind (bind (f >>> pure) >>> φ)) (isNatural _) ≡⟨ cong (λ φ bind (bind (f >>> pure) >>> φ)) (isNatural _)
bind (bind (f >>> pure) >>> 𝟙) bind (bind (f >>> pure) >>> identity)
≡⟨ cong bind .leftIdentity ≡⟨ cong bind .leftIdentity
bind (bind (f >>> pure)) bind (bind (f >>> pure))
≡⟨ cong bind (sym .rightIdentity) ≡⟨ cong bind (sym .rightIdentity)
bind (𝟙 >>> bind (f >>> pure)) ≡⟨⟩ bind (identity >>> bind (f >>> pure)) ≡⟨⟩
bind (𝟙 >=> (f >>> pure)) bind (identity >=> (f >>> pure))
≡⟨ sym (isDistributive _ _) ≡⟨ sym (isDistributive _ _)
bind 𝟙 >>> bind (f >>> pure) ≡⟨⟩ bind identity >>> bind (f >>> pure) ≡⟨⟩
bind 𝟙 >>> fmap f ≡⟨⟩ bind identity >>> fmap f ≡⟨⟩
bind 𝟙 >>> R.fmap f ≡⟨⟩ bind identity >>> R.fmap f ≡⟨⟩
R.fmap f bind 𝟙 ≡⟨⟩ R.fmap f bind identity ≡⟨⟩
R.fmap f join R.fmap f join
pureNT : NaturalTransformation R⁰ R pureNT : NaturalTransformation R⁰ R
@ -173,20 +175,20 @@ record IsMonad (raw : RawMonad) : Set where
isNaturalForeign : IsNaturalForeign isNaturalForeign : IsNaturalForeign
isNaturalForeign = begin isNaturalForeign = begin
fmap join >>> join ≡⟨⟩ fmap join >>> join ≡⟨⟩
bind (join >>> pure) >>> bind 𝟙 bind (join >>> pure) >>> bind identity
≡⟨ isDistributive _ _ ≡⟨ isDistributive _ _
bind ((join >>> pure) >>> bind 𝟙) bind ((join >>> pure) >>> bind identity)
≡⟨ cong bind .isAssociative ≡⟨ cong bind .isAssociative
bind (join >>> (pure >>> bind 𝟙)) bind (join >>> (pure >>> bind identity))
≡⟨ cong (λ φ bind (join >>> φ)) (isNatural _) ≡⟨ cong (λ φ bind (join >>> φ)) (isNatural _)
bind (join >>> 𝟙) bind (join >>> identity)
≡⟨ cong bind .leftIdentity ≡⟨ cong bind .leftIdentity
bind join ≡⟨⟩ bind join ≡⟨⟩
bind (bind 𝟙) bind (bind identity)
≡⟨ cong bind (sym .rightIdentity) ≡⟨ cong bind (sym .rightIdentity)
bind (𝟙 >>> bind 𝟙) ≡⟨⟩ bind (identity >>> bind identity) ≡⟨⟩
bind (𝟙 >=> 𝟙) ≡⟨ sym (isDistributive _ _) bind (identity >=> identity) ≡⟨ sym (isDistributive _ _)
bind 𝟙 >>> bind 𝟙 ≡⟨⟩ bind identity >>> bind identity ≡⟨⟩
join >>> join join >>> join
isInverse : IsInverse isInverse : IsInverse
@ -194,21 +196,21 @@ record IsMonad (raw : RawMonad) : Set where
where where
inv-l = begin inv-l = begin
pure >>> join ≡⟨⟩ pure >>> join ≡⟨⟩
pure >>> bind 𝟙 ≡⟨ isNatural _ pure >>> bind identity ≡⟨ isNatural _
𝟙 identity
inv-r = begin inv-r = begin
fmap pure >>> join ≡⟨⟩ fmap pure >>> join ≡⟨⟩
bind (pure >>> pure) >>> bind 𝟙 bind (pure >>> pure) >>> bind identity
≡⟨ isDistributive _ _ ≡⟨ isDistributive _ _
bind ((pure >>> pure) >=> 𝟙) ≡⟨⟩ bind ((pure >>> pure) >=> identity) ≡⟨⟩
bind ((pure >>> pure) >>> bind 𝟙) bind ((pure >>> pure) >>> bind identity)
≡⟨ cong bind .isAssociative ≡⟨ cong bind .isAssociative
bind (pure >>> (pure >>> bind 𝟙)) bind (pure >>> (pure >>> bind identity))
≡⟨ cong (λ φ bind (pure >>> φ)) (isNatural _) ≡⟨ cong (λ φ bind (pure >>> φ)) (isNatural _)
bind (pure >>> 𝟙) bind (pure >>> identity)
≡⟨ cong bind .leftIdentity ≡⟨ cong bind .leftIdentity
bind pure ≡⟨ isIdentity bind pure ≡⟨ isIdentity
𝟙 identity
record Monad : Set where record Monad : Set where
field field

View file

@ -16,8 +16,10 @@ module Cat.Category.Monad.Monoidal {a b : Level} ( : Category a b
private private
= a b = a b
open Category using (Object ; Arrow ; 𝟙 ; _∘_) open Category using (Object ; Arrow ; identity ; _∘_)
open import Cat.Category.NaturalTransformation open import Cat.Category.NaturalTransformation
using (NaturalTransformation ; Transformation ; Natural)
record RawMonad : Set where record RawMonad : Set where
field field
R : EndoFunctor R : EndoFunctor
@ -47,8 +49,8 @@ record RawMonad : Set where
joinT X Rfmap (joinT X) joinT X joinT (Romap X) joinT X Rfmap (joinT X) joinT X joinT (Romap X)
IsInverse : Set _ IsInverse : Set _
IsInverse = {X : Object} IsInverse = {X : Object}
joinT X pureT (Romap X) 𝟙 joinT X pureT (Romap X) identity
× joinT X Rfmap (pureT X) 𝟙 × joinT X Rfmap (pureT X) identity
IsNatural = {X Y} f joinT Y Rfmap f pureT X f IsNatural = {X Y} f joinT Y Rfmap f pureT X f
IsDistributive = {X Y Z} (g : Arrow Y (Romap Z)) (f : Arrow X (Romap Y)) IsDistributive = {X Y Z} (g : Arrow Y (Romap Z)) (f : Arrow X (Romap Y))
joinT Z Rfmap g (joinT Y Rfmap f) joinT Z Rfmap g (joinT Y Rfmap f)
@ -70,7 +72,7 @@ record IsMonad (raw : RawMonad) : Set where
joinT Y (R.fmap f pureT X) ≡⟨ cong (λ φ joinT Y φ) (sym (pureN f)) joinT Y (R.fmap f pureT X) ≡⟨ cong (λ φ joinT Y φ) (sym (pureN f))
joinT Y (pureT (R.omap Y) f) ≡⟨ .isAssociative joinT Y (pureT (R.omap Y) f) ≡⟨ .isAssociative
joinT Y pureT (R.omap Y) f ≡⟨ cong (λ φ φ f) (proj₁ isInverse) joinT Y pureT (R.omap Y) f ≡⟨ cong (λ φ φ f) (proj₁ isInverse)
𝟙 f ≡⟨ .leftIdentity identity f ≡⟨ .leftIdentity
f f
isDistributive : IsDistributive isDistributive : IsDistributive

View file

@ -31,7 +31,7 @@ module Cat.Category.NaturalTransformation
{c c' d d' : Level} {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 𝔻
@ -63,14 +63,14 @@ module _ (F G : Functor 𝔻) where
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 = 𝔻.identity
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 𝔻 [ 𝔻.identity F→ f ] ≡⟨ 𝔻.leftIdentity
F→ f ≡⟨ sym 𝔻.rightIdentity F→ f ≡⟨ sym 𝔻.rightIdentity
𝔻 [ F→ f 𝟙 𝔻 ] ≡⟨⟩ 𝔻 [ F→ f 𝔻.identity ] ≡⟨⟩
𝔻 [ F→ f identityTrans F A ] 𝔻 [ F→ f identityTrans F A ]
where where
module F = Functor F module F = Functor F

View file

@ -134,7 +134,7 @@ module Try0 {a b : Level} { : Category a b}
( [ ya xy ] xa) ( [ ya xy ] xa)
× [ yb xy ] xb × [ yb xy ] xb
} }
; 𝟙 = λ{ {A , f , g} .𝟙 {A} , .rightIdentity , .rightIdentity} ; identity = λ{ {A , f , g} .identity {A} , .rightIdentity , .rightIdentity}
; _∘_ = λ { {A , a0 , a1} {B , b0 , b1} {C , c0 , c1} (f , f0 , f1) (g , g0 , g1) ; _∘_ = λ { {A , a0 , a1} {B , b0 , b1} {C , c0 , c1} (f , f0 , f1) (g , g0 , g1)
(f .∘ g) (f .∘ g)
, (begin , (begin
@ -169,24 +169,24 @@ module Try0 {a b : Level} { : Category a b}
s0 = .isAssociative {f = f} {g} {h} s0 = .isAssociative {f = f} {g} {h}
isIdentity : IsIdentity 𝟙 isIdentity : IsIdentity identity
isIdentity {AA@(A , a0 , a1)} {BB@(B , b0 , b1)} {f , f0 , f1} = leftIdentity , rightIdentity isIdentity {AA@(A , a0 , a1)} {BB@(B , b0 , b1)} {f , f0 , f1} = leftIdentity , rightIdentity
where where
leftIdentity : 𝟙 (f , f0 , f1) (f , f0 , f1) leftIdentity : identity (f , f0 , f1) (f , f0 , f1)
leftIdentity i = l i , lemPropF propEqs l {proj₂ L} {proj₂ R} i leftIdentity i = l i , lemPropF propEqs l {proj₂ L} {proj₂ R} i
where where
L = 𝟙 (f , f0 , f1) L = identity (f , f0 , f1)
R : Arrow AA BB R : Arrow AA BB
R = f , f0 , f1 R = f , f0 , f1
l : proj₁ L proj₁ R l : proj₁ L proj₁ R
l = .leftIdentity l = .leftIdentity
rightIdentity : (f , f0 , f1) 𝟙 (f , f0 , f1) rightIdentity : (f , f0 , f1) identity (f , f0 , f1)
rightIdentity i = l i , lemPropF propEqs l {proj₂ L} {proj₂ R} i rightIdentity i = l i , lemPropF propEqs l {proj₂ L} {proj₂ R} i
where where
L = (f , f0 , f1) 𝟙 L = (f , f0 , f1) identity
R : Arrow AA BB R : Arrow AA BB
R = (f , f0 , f1) R = (f , f0 , f1)
l : [ f .𝟙 ] f l : [ f .identity ] f
l = .rightIdentity l = .rightIdentity
arrowsAreSets : ArrowsAreSets arrowsAreSets : ArrowsAreSets

View file

@ -52,7 +52,7 @@ module _ { : Level} { : Category } where
isIdentity : IsIdentity isIdentity : IsIdentity
isIdentity {c} = lemSig prp _ _ eq isIdentity {c} = lemSig prp _ _ eq
where where
eq : (λ C x [ .𝟙 x ]) identityTrans (presheaf c) eq : (λ C x [ .identity x ]) identityTrans (presheaf c)
eq = funExt λ A funExt λ B .leftIdentity eq = funExt λ A funExt λ B .leftIdentity
prp = F.naturalIsProp _ _ {F = presheaf c} {presheaf c} prp = F.naturalIsProp _ _ {F = presheaf c} {presheaf c}