Rename func* and func-> to omap and fmap respectively

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-08 11:03:56 +01:00
parent 2fcc583646
commit 5ad506a09f
7 changed files with 142 additions and 143 deletions

View file

@ -17,7 +17,7 @@ open import Cat.Category.NaturalTransformation
open import Cat.Equality open import Cat.Equality
open Equality.Data.Product open Equality.Data.Product
open Functor using (func→ ; func*) open Functor using (fmap ; omap)
open Category using (Object ; 𝟙) open Category using (Object ; 𝟙)
-- The category of categories -- The category of categories
@ -104,13 +104,13 @@ module CatProduct { ' : Level} ( 𝔻 : Category ') where
proj₁ : Functor obj proj₁ : Functor obj
proj₁ = record proj₁ = record
{ raw = record { func* = fst ; func→ = fst } { raw = record { omap = fst ; fmap = fst }
; isFunctor = record { isIdentity = refl ; isDistributive = refl } ; isFunctor = record { isIdentity = refl ; isDistributive = refl }
} }
proj₂ : Functor obj 𝔻 proj₂ : Functor obj 𝔻
proj₂ = record proj₂ = record
{ raw = record { func* = snd ; func→ = snd } { raw = record { omap = snd ; fmap = snd }
; isFunctor = record { isIdentity = refl ; isDistributive = refl } ; isFunctor = record { isIdentity = refl ; isDistributive = refl }
} }
@ -119,8 +119,8 @@ module CatProduct { ' : Level} ( 𝔻 : Category ') where
x : Functor X obj x : Functor X obj
x = record x = record
{ raw = record { raw = record
{ func* = λ x x₁.func* x , x₂.func* x { omap = λ x x₁.omap x , x₂.omap x
; func→ = λ x x₁.func→ x , x₂.func→ x ; fmap = λ x x₁.fmap x , x₂.fmap x
} }
; isFunctor = record ; isFunctor = record
{ isIdentity = Σ≡ x₁.isIdentity x₂.isIdentity { isIdentity = Σ≡ x₁.isIdentity x₂.isIdentity
@ -175,8 +175,8 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
Category = Category Category = Category
open Fun 𝔻 renaming (identity to idN) open Fun 𝔻 renaming (identity to idN)
private private
:func*: : Functor 𝔻 × Object Object 𝔻 :omap: : Functor 𝔻 × Object Object 𝔻
:func*: (F , A) = func* F A :omap: (F , A) = omap F A
prodObj : Category prodObj : Category
prodObj = Fun prodObj = Fun
@ -193,28 +193,28 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
B : Object B : Object
B = proj₂ cod B = proj₂ cod
:func→: : (pobj : NaturalTransformation F G × [ A , B ]) :fmap: : (pobj : NaturalTransformation F G × [ A , B ])
𝔻 [ func* F A , func* G B ] 𝔻 [ omap F A , omap G B ]
:func→: ((θ , θNat) , f) = result :fmap: ((θ , θNat) , f) = result
where where
θA : 𝔻 [ func* F A , func* G A ] θA : 𝔻 [ omap F A , omap G A ]
θA = θ A θA = θ A
θB : 𝔻 [ func* F B , func* G B ] θB : 𝔻 [ omap F B , omap G B ]
θB = θ B θB = θ B
F→f : 𝔻 [ func* F A , func* F B ] F→f : 𝔻 [ omap F A , omap F B ]
F→f = func→ F f F→f = fmap F f
G→f : 𝔻 [ func* G A , func* G B ] G→f : 𝔻 [ omap G A , omap G B ]
G→f = func→ G f G→f = fmap G f
l : 𝔻 [ func* F A , func* G B ] l : 𝔻 [ omap F A , omap G B ]
l = 𝔻 [ θB F→f ] l = 𝔻 [ θB F→f ]
r : 𝔻 [ func* F A , func* G B ] r : 𝔻 [ omap F A , omap G B ]
r = 𝔻 [ G→f θA ] r = 𝔻 [ G→f θA ]
-- There are two choices at this point, -- There are two choices at this point,
-- but I suppose the whole point is that -- but I suppose the whole point is that
-- by `θNat f` we have `l ≡ r` -- by `θNat f` we have `l ≡ r`
-- lem : 𝔻 [ θ B ∘ F .func→ f ] ≡ 𝔻 [ G .func→ f ∘ θ A ] -- lem : 𝔻 [ θ B ∘ F .fmap f ] ≡ 𝔻 [ G .fmap f ∘ θ A ]
-- lem = θNat f -- lem = θNat f
result : 𝔻 [ func* F A , func* G B ] result : 𝔻 [ omap F A , omap G B ]
result = l result = l
open CatProduct renaming (obj to _×p_) using () open CatProduct renaming (obj to _×p_) using ()
@ -227,19 +227,19 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
C = proj₂ c C = proj₂ c
-- NaturalTransformation F G × .Arrow A B -- NaturalTransformation F G × .Arrow A B
-- :ident: : :func→: {c} {c} (identityNat F , .𝟙) 𝔻 .𝟙 -- :ident: : :fmap: {c} {c} (identityNat F , .𝟙) 𝔻 .𝟙
-- :ident: = trans (proj₂ 𝔻.isIdentity) (F .isIdentity) -- :ident: = trans (proj₂ 𝔻.isIdentity) (F .isIdentity)
-- where -- where
-- open module 𝔻 = IsCategory (𝔻 .isCategory) -- open module 𝔻 = IsCategory (𝔻 .isCategory)
-- Unfortunately the equational version has some ambigous arguments. -- Unfortunately the equational version has some ambigous arguments.
:ident: : :func→: {c} {c} (NT.identity F , 𝟙 {A = proj₂ c}) 𝟙 𝔻 :ident: : :fmap: {c} {c} (NT.identity F , 𝟙 {A = proj₂ c}) 𝟙 𝔻
:ident: = begin :ident: = begin
:func→: {c} {c} (𝟙 (prodObj ×p ) {c}) ≡⟨⟩ :fmap: {c} {c} (𝟙 (prodObj ×p ) {c}) ≡⟨⟩
:func→: {c} {c} (idN F , 𝟙 ) ≡⟨⟩ :fmap: {c} {c} (idN F , 𝟙 ) ≡⟨⟩
𝔻 [ identityTrans F C func→ F (𝟙 )] ≡⟨⟩ 𝔻 [ identityTrans F C fmap F (𝟙 )] ≡⟨⟩
𝔻 [ 𝟙 𝔻 func→ F (𝟙 )] ≡⟨ proj₂ 𝔻.isIdentity 𝔻 [ 𝟙 𝔻 fmap F (𝟙 )] ≡⟨ proj₂ 𝔻.isIdentity
func→ F (𝟙 ) ≡⟨ F.isIdentity fmap F (𝟙 ) ≡⟨ F.isIdentity
𝟙 𝔻 𝟙 𝔻
where where
open module 𝔻 = Category 𝔻 open module 𝔻 = Category 𝔻
@ -279,28 +279,28 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
ηθNat = proj₂ ηθNT ηθNat = proj₂ ηθNT
:isDistributive: : :isDistributive: :
𝔻 [ 𝔻 [ η C θ C ] func→ F ( [ g f ] ) ] 𝔻 [ 𝔻 [ η C θ C ] fmap F ( [ g f ] ) ]
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ θ B func→ F f ] ] 𝔻 [ 𝔻 [ η C fmap G g ] 𝔻 [ θ B fmap F f ] ]
:isDistributive: = begin :isDistributive: = begin
𝔻 [ (ηθ C) func→ F ( [ g f ]) ] 𝔻 [ (ηθ C) fmap F ( [ g f ]) ]
≡⟨ ηθNat ( [ g f ]) ≡⟨ ηθNat ( [ g f ])
𝔻 [ func→ H ( [ g f ]) (ηθ A) ] 𝔻 [ fmap H ( [ g f ]) (ηθ A) ]
≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.isDistributive) ≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.isDistributive)
𝔻 [ 𝔻 [ func→ H g func→ H f ] (ηθ A) ] 𝔻 [ 𝔻 [ fmap H g fmap H f ] (ηθ A) ]
≡⟨ sym isAssociative ≡⟨ sym isAssociative
𝔻 [ func→ H g 𝔻 [ func→ H f ηθ A ] ] 𝔻 [ fmap H g 𝔻 [ fmap H f ηθ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) isAssociative ≡⟨ cong (λ φ 𝔻 [ fmap H g φ ]) isAssociative
𝔻 [ func→ H g 𝔻 [ 𝔻 [ func→ H f η A ] θ A ] ] 𝔻 [ fmap H g 𝔻 [ 𝔻 [ fmap H f η A ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f))) ≡⟨ cong (λ φ 𝔻 [ fmap H g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f)))
𝔻 [ func→ H g 𝔻 [ 𝔻 [ η B func→ G f ] θ A ] ] 𝔻 [ fmap H g 𝔻 [ 𝔻 [ η B fmap G f ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ func→ H g φ ]) (sym isAssociative) ≡⟨ cong (λ φ 𝔻 [ fmap H g φ ]) (sym isAssociative)
𝔻 [ func→ H g 𝔻 [ η B 𝔻 [ func→ G f θ A ] ] ] 𝔻 [ fmap H g 𝔻 [ η B 𝔻 [ fmap G f θ A ] ] ]
≡⟨ isAssociative ≡⟨ isAssociative
𝔻 [ 𝔻 [ func→ H g η B ] 𝔻 [ func→ G f θ A ] ] 𝔻 [ 𝔻 [ fmap H g η B ] 𝔻 [ fmap G f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ func→ G f θ A ] ]) (sym (ηNat g)) ≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ fmap G f θ A ] ]) (sym (ηNat g))
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ func→ G f θ A ] ] 𝔻 [ 𝔻 [ η C fmap G g ] 𝔻 [ fmap G f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C func→ G g ] φ ]) (sym (θNat f)) ≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C fmap G g ] φ ]) (sym (θNat f))
𝔻 [ 𝔻 [ η C func→ G g ] 𝔻 [ θ B func→ F f ] ] 𝔻 [ 𝔻 [ η C fmap G g ] 𝔻 [ θ B fmap F f ] ]
where where
open Category 𝔻 open Category 𝔻
module H = Functor H module H = Functor H
@ -309,8 +309,8 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
-- :eval: : Functor (prodObj ×p ) 𝔻 -- :eval: : Functor (prodObj ×p ) 𝔻
eval = record eval = record
{ raw = record { raw = record
{ func* = :func*: { omap = :omap:
; func→ = λ {dom} {cod} :func→: {dom} {cod} ; fmap = λ {dom} {cod} :fmap: {dom} {cod}
} }
; isFunctor = record ; isFunctor = record
{ isIdentity = λ {o} :ident: {o} { isIdentity = λ {o} :ident: {o}

View file

@ -28,20 +28,20 @@ module _ {a b : Level} where
private private
module T = Functor T module T = Functor T
Type : (Γ : Object ) Set a Type : (Γ : Object ) Set a
Type Γ = proj₁ (proj₁ (T.func* Γ)) Type Γ = proj₁ (proj₁ (T.omap Γ))
module _ {Γ : Object } {A : Type Γ} where module _ {Γ : Object } {A : Type Γ} where
-- module _ {A B : Object } {γ : [ A , B ]} where -- module _ {A B : Object } {γ : [ A , B ]} where
-- k : Σ (proj₁ (func* T B) → proj₁ (func* T A)) -- k : Σ (proj₁ (omap T B) → proj₁ (omap T A))
-- (λ f → -- (λ f →
-- {x : proj₁ (func* T B)} → -- {x : proj₁ (omap T B)} →
-- proj₂ (func* T B) x → proj₂ (func* T A) (f x)) -- proj₂ (omap T B) x → proj₂ (omap T A) (f x))
-- k = T.func→ γ -- k = T.fmap γ
-- k₁ : proj₁ (func* T B) → proj₁ (func* T A) -- k₁ : proj₁ (omap T B) → proj₁ (omap T A)
-- k₁ = proj₁ k -- k₁ = proj₁ k
-- k₂ : ({x : proj₁ (func* T B)} → -- k₂ : ({x : proj₁ (omap T B)} →
-- proj₂ (func* T B) x → proj₂ (func* T A) (k₁ x)) -- proj₂ (omap T B) x → proj₂ (omap T A) (k₁ x))
-- k₂ = proj₂ k -- k₂ = proj₂ k
record ContextComprehension : Set (a b) where record ContextComprehension : Set (a b) where
@ -51,7 +51,7 @@ module _ {a b : Level} where
-- proj2 : ???? -- proj2 : ????
-- if γ : [ A , B ] -- if γ : [ A , B ]
-- then T .func→ γ (written T[γ]) interpret substitutions in types and terms respectively. -- then T .fmap γ (written T[γ]) interpret substitutions in types and terms respectively.
-- field -- field
-- ump : {Δ : .Object} → (γ : [ Δ , Γ ]) -- ump : {Δ : .Object} → (γ : [ Δ , Γ ])
-- → (a : {!!}) → {!!} -- → (a : {!!}) → {!!}

View file

@ -97,8 +97,8 @@ module _ {a b : Level} ( : Category a b) where
representable : Category.Object Representable representable : Category.Object Representable
representable A = record representable A = record
{ raw = record { raw = record
{ func* = λ B [ A , B ] , arrowsAreSets { omap = λ B [ A , B ] , arrowsAreSets
; func→ = [_∘_] ; fmap = [_∘_]
} }
; isFunctor = record ; isFunctor = record
{ isIdentity = funExt λ _ proj₂ isIdentity { isIdentity = funExt λ _ proj₂ isIdentity
@ -110,8 +110,8 @@ module _ {a b : Level} ( : Category a b) where
presheaf : Category.Object (opposite ) Presheaf presheaf : Category.Object (opposite ) Presheaf
presheaf B = record presheaf B = record
{ raw = record { raw = record
{ func* = λ A [ A , B ] , arrowsAreSets { omap = λ A [ A , B ] , arrowsAreSets
; func→ = λ f g [ g f ] ; fmap = λ f g [ g f ]
} }
; isFunctor = record ; isFunctor = record
{ isIdentity = funExt λ x proj₁ isIdentity { isIdentity = funExt λ x proj₁ isIdentity

View file

@ -24,39 +24,39 @@ module _ {c c' d d'}
[ A , B ] 𝔻 [ omap A , omap B ] [ A , B ] 𝔻 [ omap A , omap B ]
record RawFunctor : 𝓤 where record RawFunctor : 𝓤 where
field field
func* : Object Object 𝔻 omap : Object Object 𝔻
func→ : {A B} [ A , B ] 𝔻 [ func* A , func* B ] fmap : {A B} [ A , B ] 𝔻 [ omap A , omap B ]
IsIdentity : Set _ IsIdentity : Set _
IsIdentity = {A : Object } func→ (𝟙 {A}) 𝟙 𝔻 {func* A} IsIdentity = {A : Object } fmap (𝟙 {A}) 𝟙 𝔻 {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 ]}
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ] fmap ( [ g f ]) 𝔻 [ fmap g fmap f ]
-- | Equality principle for raw functors -- | Equality principle for raw functors
-- --
-- The type of `func→` depend on the value of `func*`. We can wrap this up -- The type of `fmap` depend on the value of `omap`. We can wrap this up
-- into an equality principle for this type like is done for e.g. `Σ` using -- into an equality principle for this type like is done for e.g. `Σ` using
-- `pathJ`. -- `pathJ`.
module _ {x y : RawFunctor} where module _ {x y : RawFunctor} where
open RawFunctor open RawFunctor
private private
P : (omap : Omap) (eq : func* x omap) Set _ P : (omap' : Omap) (eq : omap x omap') Set _
P y eq = (fmap' : Fmap y) (λ i Fmap (eq i)) P y eq = (fmap' : Fmap y) (λ i Fmap (eq i))
[ func→ x fmap' ] [ fmap x fmap' ]
module _ module _
(eq : (λ i Omap) [ func* x func* y ]) (eq : (λ i Omap) [ omap x omap y ])
(kk : P (func* x) refl) (kk : P (omap x) refl)
where where
private private
p : P (func* y) eq p : P (omap y) eq
p = pathJ P kk (func* y) eq p = pathJ P kk (omap y) eq
eq→ : (λ i Fmap (eq i)) [ func→ x func→ y ] eq→ : (λ i Fmap (eq i)) [ fmap x fmap y ]
eq→ = p (func→ y) eq→ = p (fmap y)
RawFunctor≡ : x y RawFunctor≡ : x y
func* (RawFunctor≡ i) = eq i omap (RawFunctor≡ i) = eq i
func→ (RawFunctor≡ i) = eq→ i fmap (RawFunctor≡ i) = eq→ i
record IsFunctor (F : RawFunctor) : 𝓤 where record IsFunctor (F : RawFunctor) : 𝓤 where
open RawFunctor F public open RawFunctor F public
@ -124,10 +124,10 @@ module _ { ' : Level} { 𝔻 : Category '} where
module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where
private private
F* = func* F F* = omap F
F→ = func→ F F→ = fmap F
G* = func* G G* = omap G
G→ = func→ G G→ = fmap G
module _ {a0 a1 a2 : Object A} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where module _ {a0 a1 a2 : Object A} {α0 : A [ a0 , a1 ]} {α1 : A [ a1 , a2 ]} where
dist : (F→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ] dist : (F→ G→) (A [ α1 α0 ]) C [ (F→ G→) α1 (F→ G→) α0 ]
@ -138,8 +138,8 @@ module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : F
C [ (F→ G→) α1 (F→ G→) α0 ] C [ (F→ G→) α1 (F→ G→) α0 ]
_∘fr_ : RawFunctor A C _∘fr_ : RawFunctor A C
RawFunctor.func* _∘fr_ = F* G* RawFunctor.omap _∘fr_ = F* G*
RawFunctor.func→ _∘fr_ = F→ G→ RawFunctor.fmap _∘fr_ = F→ G→
instance instance
isFunctor' : IsFunctor A C _∘fr_ isFunctor' : IsFunctor A C _∘fr_
isFunctor' = record isFunctor' = record
@ -158,8 +158,8 @@ module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : F
identity : { '} {C : Category '} Functor C C identity : { '} {C : Category '} Functor C C
identity = record identity = record
{ raw = record { raw = record
{ func* = λ x x { omap = λ x x
; func→ = λ x x ; fmap = λ x x
} }
; isFunctor = record ; isFunctor = record
{ isIdentity = refl { isIdentity = refl

View file

@ -39,8 +39,8 @@ module Monoidal {a b : Level} ( : Category a b) where
joinN : Natural F[ R R ] R joinT joinN : Natural F[ R R ] R joinT
joinN = proj₂ joinNT joinN = proj₂ joinNT
Romap = Functor.func* R Romap = Functor.omap R
Rfmap = Functor.func→ R Rfmap = Functor.fmap R
bind : {X Y : Object} [ X , Romap Y ] [ Romap X , Romap Y ] bind : {X Y : Object} [ X , Romap Y ] [ Romap X , Romap Y ]
bind {X} {Y} f = joinT Y Rfmap f bind {X} {Y} f = joinT Y Rfmap f
@ -69,10 +69,10 @@ module Monoidal {a b : Level} ( : Category a b) where
isNatural : IsNatural isNatural : IsNatural
isNatural {X} {Y} f = begin isNatural {X} {Y} f = begin
joinT Y R.func→ f pureT X ≡⟨ sym .isAssociative joinT Y R.fmap f pureT X ≡⟨ sym .isAssociative
joinT Y (R.func→ 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.func* Y) f) ≡⟨ .isAssociative joinT Y (pureT (R.omap Y) f) ≡⟨ .isAssociative
joinT Y pureT (R.func* Y) f ≡⟨ cong (λ φ φ f) (proj₁ isInverse) joinT Y pureT (R.omap Y) f ≡⟨ cong (λ φ φ f) (proj₁ isInverse)
𝟙 f ≡⟨ proj₂ .isIdentity 𝟙 f ≡⟨ proj₂ .isIdentity
f f
@ -81,36 +81,36 @@ module Monoidal {a b : Level} ( : Category a b) where
where where
module R² = Functor F[ R R ] module R² = Functor F[ R R ]
distrib3 : {A B C D} {a : Arrow C D} {b : Arrow B C} {c : Arrow A B} distrib3 : {A B C D} {a : Arrow C D} {b : Arrow B C} {c : Arrow A B}
R.func→ (a b c) R.fmap (a b c)
R.func→ a R.func→ b R.func→ c R.fmap a R.fmap b R.fmap c
distrib3 {a = a} {b} {c} = begin distrib3 {a = a} {b} {c} = begin
R.func→ (a b c) ≡⟨ R.isDistributive R.fmap (a b c) ≡⟨ R.isDistributive
R.func→ (a b) R.func→ c ≡⟨ cong (_∘ _) R.isDistributive R.fmap (a b) R.fmap c ≡⟨ cong (_∘ _) R.isDistributive
R.func→ a R.func→ b R.func→ c R.fmap a R.fmap b R.fmap c
aux = begin aux = begin
joinT Z R.func→ (joinT Z R.func→ g f) joinT Z R.fmap (joinT Z R.fmap g f)
≡⟨ cong (λ φ joinT Z φ) distrib3 ≡⟨ cong (λ φ joinT Z φ) distrib3
joinT Z (R.func→ (joinT Z) R.func→ (R.func→ g) R.func→ f) joinT Z (R.fmap (joinT Z) R.fmap (R.fmap g) R.fmap f)
≡⟨⟩ ≡⟨⟩
joinT Z (R.func→ (joinT Z) R².func→ g R.func→ f) joinT Z (R.fmap (joinT Z) R².fmap g R.fmap f)
≡⟨ cong (_∘_ (joinT Z)) (sym .isAssociative) ≡⟨ cong (_∘_ (joinT Z)) (sym .isAssociative)
joinT Z (R.func→ (joinT Z) (R².func→ g R.func→ f)) joinT Z (R.fmap (joinT Z) (R².fmap g R.fmap f))
≡⟨ .isAssociative ≡⟨ .isAssociative
(joinT Z R.func→ (joinT Z)) (R².func→ g R.func→ f) (joinT Z R.fmap (joinT Z)) (R².fmap g R.fmap f)
≡⟨ cong (λ φ φ (R².func→ g R.func→ f)) isAssociative ≡⟨ cong (λ φ φ (R².fmap g R.fmap f)) isAssociative
(joinT Z joinT (R.func* Z)) (R².func→ g R.func→ f) (joinT Z joinT (R.omap Z)) (R².fmap g R.fmap f)
≡⟨ .isAssociative ≡⟨ .isAssociative
joinT Z joinT (R.func* Z) R².func→ g R.func→ f joinT Z joinT (R.omap Z) R².fmap g R.fmap f
≡⟨⟩ ≡⟨⟩
((joinT Z joinT (R.func* Z)) R².func→ g) R.func→ f ((joinT Z joinT (R.omap Z)) R².fmap g) R.fmap f
≡⟨ cong (_∘ R.func→ f) (sym .isAssociative) ≡⟨ cong (_∘ R.fmap f) (sym .isAssociative)
(joinT Z (joinT (R.func* Z) R².func→ g)) R.func→ f (joinT Z (joinT (R.omap Z) R².fmap g)) R.fmap f
≡⟨ cong (λ φ φ R.func→ f) (cong (_∘_ (joinT Z)) (joinN g)) ≡⟨ cong (λ φ φ R.fmap f) (cong (_∘_ (joinT Z)) (joinN g))
(joinT Z (R.func→ g joinT Y)) R.func→ f (joinT Z (R.fmap g joinT Y)) R.fmap f
≡⟨ cong (_∘ R.func→ f) .isAssociative ≡⟨ cong (_∘ R.fmap f) .isAssociative
joinT Z R.func→ g joinT Y R.func→ f joinT Z R.fmap g joinT Y R.fmap f
≡⟨ sym (Category.isAssociative ) ≡⟨ sym (Category.isAssociative )
joinT Z R.func→ g (joinT Y R.func→ f) joinT Z R.fmap g (joinT Y R.fmap f)
record Monad : Set where record Monad : Set where
@ -235,8 +235,8 @@ module Kleisli {a b : Level} ( : Category a b) where
-- | This formulation gives rise to the following endo-functor. -- | This formulation gives rise to the following endo-functor.
private private
rawR : RawFunctor rawR : RawFunctor
RawFunctor.func* rawR = omap RawFunctor.omap rawR = omap
RawFunctor.func→ rawR = fmap RawFunctor.fmap rawR = fmap
isFunctorR : IsFunctor rawR isFunctorR : IsFunctor rawR
IsFunctor.isIdentity isFunctorR = begin IsFunctor.isIdentity isFunctorR = begin
@ -269,18 +269,18 @@ module Kleisli {a b : Level} ( : Category a b) where
pureT A = pure pureT A = pure
pureN : Natural R⁰ R pureT pureN : Natural R⁰ R pureT
pureN {A} {B} f = begin pureN {A} {B} f = begin
pureT B R⁰.func→ f ≡⟨⟩ pureT B R⁰.fmap f ≡⟨⟩
pure f ≡⟨ sym (isNatural _) pure f ≡⟨ sym (isNatural _)
bind (pure f) pure ≡⟨⟩ bind (pure f) pure ≡⟨⟩
fmap f pure ≡⟨⟩ fmap f pure ≡⟨⟩
R.func→ f pureT A R.fmap f pureT A
joinT : Transformation R joinT : Transformation R
joinT C = join joinT C = join
joinN : Natural R joinT joinN : Natural R joinT
joinN f = begin joinN f = begin
join R².func→ f ≡⟨⟩ join R².fmap f ≡⟨⟩
bind 𝟙 R².func→ f ≡⟨⟩ bind 𝟙 R².fmap f ≡⟨⟩
R².func→ f >>> bind 𝟙 ≡⟨⟩ R².fmap f >>> bind 𝟙 ≡⟨⟩
fmap (fmap f) >>> bind 𝟙 ≡⟨⟩ fmap (fmap f) >>> bind 𝟙 ≡⟨⟩
fmap (bind (f >>> pure)) >>> bind 𝟙 ≡⟨⟩ fmap (bind (f >>> pure)) >>> bind 𝟙 ≡⟨⟩
bind (bind (f >>> pure) >>> pure) >>> bind 𝟙 bind (bind (f >>> pure) >>> pure) >>> bind 𝟙
@ -300,9 +300,9 @@ module Kleisli {a b : Level} ( : Category a b) where
≡⟨ sym (isDistributive _ _) ≡⟨ sym (isDistributive _ _)
bind 𝟙 >>> bind (f >>> pure) ≡⟨⟩ bind 𝟙 >>> bind (f >>> pure) ≡⟨⟩
bind 𝟙 >>> fmap f ≡⟨⟩ bind 𝟙 >>> fmap f ≡⟨⟩
bind 𝟙 >>> R.func→ f ≡⟨⟩ bind 𝟙 >>> R.fmap f ≡⟨⟩
R.func→ f bind 𝟙 ≡⟨⟩ R.fmap f bind 𝟙 ≡⟨⟩
R.func→ f join R.fmap f join
pureNT : NaturalTransformation R⁰ R pureNT : NaturalTransformation R⁰ R
proj₁ pureNT = pureT proj₁ pureNT = pureT
@ -396,7 +396,6 @@ module _ {a b : Level} { : Category a b} where
private private
module = Category module = Category
open using (Object ; Arrow ; 𝟙 ; _∘_ ; _>>>_) open using (Object ; Arrow ; 𝟙 ; _∘_ ; _>>>_)
open Functor using (func* ; func→)
module M = Monoidal module M = Monoidal
module K = Kleisli module K = Kleisli
@ -434,19 +433,19 @@ module _ {a b : Level} { : Category a b} where
backIsMonad : M.IsMonad backRaw backIsMonad : M.IsMonad backRaw
M.IsMonad.isAssociative backIsMonad {X} = begin M.IsMonad.isAssociative backIsMonad {X} = begin
joinT X R.func→ (joinT X) ≡⟨⟩ joinT X R.fmap (joinT X) ≡⟨⟩
join fmap (joinT X) ≡⟨⟩ join fmap (joinT X) ≡⟨⟩
join fmap join ≡⟨ isNaturalForeign join fmap join ≡⟨ isNaturalForeign
join join ≡⟨⟩ join join ≡⟨⟩
joinT X joinT (R.func* X) joinT X joinT (R.omap X)
M.IsMonad.isInverse backIsMonad {X} = inv-l , inv-r M.IsMonad.isInverse backIsMonad {X} = inv-l , inv-r
where where
inv-l = begin inv-l = begin
joinT X pureT (R.func* X) ≡⟨⟩ joinT X pureT (R.omap X) ≡⟨⟩
join pure ≡⟨ proj₁ isInverse join pure ≡⟨ proj₁ isInverse
𝟙 𝟙
inv-r = begin inv-r = begin
joinT X R.func→ (pureT X) ≡⟨⟩ joinT X R.fmap (pureT X) ≡⟨⟩
join fmap pure ≡⟨ proj₂ isInverse join fmap pure ≡⟨ proj₂ isInverse
𝟙 𝟙
@ -526,8 +525,8 @@ module _ {a b : Level} { : Category a b} where
) )
rawEq : Functor.raw KM.R Functor.raw R rawEq : Functor.raw KM.R Functor.raw R
RawFunctor.func* (rawEq i) = omapEq i RawFunctor.omap (rawEq i) = omapEq i
RawFunctor.func→ (rawEq i) = fmapEq i RawFunctor.fmap (rawEq i) = fmapEq i
Req : M.RawMonad.R (backRaw (forth m)) R Req : M.RawMonad.R (backRaw (forth m)) R
Req = Functor≡ rawEq Req = Functor≡ rawEq
@ -586,8 +585,8 @@ module _ {a b : Level} ( : Category a b) where
Rraw : RawFunctor Rraw : RawFunctor
Rraw = record Rraw = record
{ func* = omap { omap = omap
; func→ = fmap ; fmap = fmap
} }
field field
@ -663,7 +662,7 @@ module _ {a b : Level} { : Category a b} where
voe-2-3-1-fromMonad : (m : M.Monad) voe-2-3-1 (M.Monad.Romap m) (λ {X} M.Monad.pureT m X) voe-2-3-1-fromMonad : (m : M.Monad) voe-2-3-1 (M.Monad.Romap m) (λ {X} M.Monad.pureT m X)
voe-2-3-1-fromMonad m = record voe-2-3-1-fromMonad m = record
{ fmap = Functor.func→ R { fmap = Functor.fmap R
; RisFunctor = Functor.isFunctor R ; RisFunctor = Functor.isFunctor R
; pureN = pureN ; pureN = pureN
; join = λ {X} joinT X ; join = λ {X} joinT X

View file

@ -46,13 +46,13 @@ module NaturalTransformation {c c' d d' : Level}
module G = Functor G module G = Functor G
-- What do you call a non-natural tranformation? -- What do you call a non-natural tranformation?
Transformation : Set (c d') Transformation : Set (c d')
Transformation = (C : Object ) 𝔻 [ F.func* C , G.func* C ] Transformation = (C : Object ) 𝔻 [ F.omap C , G.omap C ]
Natural : Transformation Set (c (c' d')) Natural : Transformation Set (c (c' d'))
Natural θ Natural θ
= {A B : Object } = {A B : Object }
(f : [ A , B ]) (f : [ A , B ])
𝔻 [ θ B F.func→ f ] 𝔻 [ G.func→ f θ A ] 𝔻 [ θ B F.fmap f ] 𝔻 [ G.fmap f θ A ]
NaturalTransformation : Set (c c' d') NaturalTransformation : Set (c c' d')
NaturalTransformation = Σ Transformation Natural NaturalTransformation = Σ Transformation Natural
@ -78,7 +78,7 @@ module NaturalTransformation {c c' d d' : Level}
𝔻 [ F→ f identityTrans F A ] 𝔻 [ F→ f identityTrans F A ]
where where
module F = Functor F module F = Functor F
F→ = F.func→ 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
@ -94,14 +94,14 @@ module NaturalTransformation {c c' d d' : Level}
NT[_∘_] : NaturalTransformation G H NaturalTransformation F G NaturalTransformation F H NT[_∘_] : NaturalTransformation G H NaturalTransformation F G NaturalTransformation F H
proj₁ NT[ (θ , _) (η , _) ] = T[ θ η ] proj₁ NT[ (θ , _) (η , _) ] = T[ θ η ]
proj₂ NT[ (θ , θNat) (η , ηNat) ] {A} {B} f = begin proj₂ NT[ (θ , θNat) (η , ηNat) ] {A} {B} f = begin
𝔻 [ T[ θ η ] B F.func→ f ] ≡⟨⟩ 𝔻 [ T[ θ η ] B F.fmap f ] ≡⟨⟩
𝔻 [ 𝔻 [ θ B η B ] F.func→ f ] ≡⟨ sym 𝔻.isAssociative 𝔻 [ 𝔻 [ θ B η B ] F.fmap f ] ≡⟨ sym 𝔻.isAssociative
𝔻 [ θ B 𝔻 [ η B F.func→ f ] ] ≡⟨ cong (λ φ 𝔻 [ θ B φ ]) (ηNat f) 𝔻 [ θ B 𝔻 [ η B F.fmap f ] ] ≡⟨ cong (λ φ 𝔻 [ θ B φ ]) (ηNat f)
𝔻 [ θ B 𝔻 [ G.func→ f η A ] ] ≡⟨ 𝔻.isAssociative 𝔻 [ θ B 𝔻 [ G.fmap f η A ] ] ≡⟨ 𝔻.isAssociative
𝔻 [ 𝔻 [ θ B G.func→ f ] η A ] ≡⟨ cong (λ φ 𝔻 [ φ η A ]) (θNat f) 𝔻 [ 𝔻 [ θ B G.fmap f ] η A ] ≡⟨ cong (λ φ 𝔻 [ φ η A ]) (θNat f)
𝔻 [ 𝔻 [ H.func→ f θ A ] η A ] ≡⟨ sym 𝔻.isAssociative 𝔻 [ 𝔻 [ H.fmap f θ A ] η A ] ≡⟨ sym 𝔻.isAssociative
𝔻 [ H.func→ f 𝔻 [ θ A η A ] ] ≡⟨⟩ 𝔻 [ H.fmap f 𝔻 [ θ A η A ] ] ≡⟨⟩
𝔻 [ H.func→ f T[ θ η ] A ] 𝔻 [ H.fmap f T[ θ η ] A ]
module _ {F G : Functor 𝔻} where module _ {F G : Functor 𝔻} where
transformationIsSet : isSet (Transformation F G) transformationIsSet : isSet (Transformation F G)

View file

@ -42,9 +42,9 @@ module _ { : Level} { : Category } where
fmapNT = fmap , fmapNatural fmapNT = fmap , fmapNatural
rawYoneda : RawFunctor Fun rawYoneda : RawFunctor Fun
RawFunctor.func* rawYoneda = prshf RawFunctor.omap rawYoneda = prshf
RawFunctor.func→ rawYoneda = fmapNT RawFunctor.fmap rawYoneda = fmapNT
open RawFunctor rawYoneda open RawFunctor rawYoneda hiding (fmap)
isIdentity : IsIdentity isIdentity : IsIdentity
isIdentity {c} = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq isIdentity {c} = lemSig (naturalIsProp {F = prshf c} {prshf c}) _ _ eq