Prove that fmap is mapped correctly

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-06 15:52:22 +01:00
parent 4d528a7077
commit 5ae68df582
3 changed files with 62 additions and 62 deletions

View file

@ -76,7 +76,7 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
𝟙 : {A : Object} Arrow A A 𝟙 : {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 _∘_ _>>>_
-- | Operations on data -- | Operations on data

View file

@ -454,30 +454,31 @@ module _ {a b : Level} { : Category a b} where
Monoidal.Monad.isMonad (back m) = backIsMonad m Monoidal.Monad.isMonad (back m) = backIsMonad m
module _ (m : K.Monad) where module _ (m : K.Monad) where
open K.Monad m private
bindEq : {X Y} open K.Monad m
K.RawMonad.bind (forthRaw (backRaw m)) {X} {Y} bindEq : {X Y}
K.RawMonad.bind (K.Monad.raw m) K.RawMonad.bind (forthRaw (backRaw m)) {X} {Y}
bindEq {X} {Y} = begin K.RawMonad.bind (K.Monad.raw m)
K.RawMonad.bind (forthRaw (backRaw m)) ≡⟨⟩ bindEq {X} {Y} = begin
(λ f join fmap f) ≡⟨⟩ K.RawMonad.bind (forthRaw (backRaw m)) ≡⟨⟩
(λ f bind (f >>> pure) >>> bind 𝟙) ≡⟨ funExt lem (λ f join fmap f) ≡⟨⟩
(λ f bind f) ≡⟨⟩ (λ f bind (f >>> pure) >>> bind 𝟙) ≡⟨ funExt lem
bind (λ f bind f) ≡⟨⟩
where bind
lem : (f : Arrow X (omap Y)) where
bind (f >>> pure) >>> bind 𝟙 lem : (f : Arrow X (omap Y))
bind f bind (f >>> pure) >>> bind 𝟙
lem f = begin bind f
bind (f >>> pure) >>> bind 𝟙 lem f = begin
≡⟨ isDistributive _ _ bind (f >>> pure) >>> bind 𝟙
bind ((f >>> pure) >>> bind 𝟙) ≡⟨ isDistributive _ _
≡⟨ cong bind .isAssociative bind ((f >>> pure) >>> bind 𝟙)
bind (f >>> (pure >>> bind 𝟙)) ≡⟨ cong bind .isAssociative
≡⟨ cong (λ φ bind (f >>> φ)) (isNatural _) bind (f >>> (pure >>> bind 𝟙))
bind (f >>> 𝟙) ≡⟨ cong (λ φ bind (f >>> φ)) (isNatural _)
≡⟨ cong bind (proj₂ .isIdentity) bind (f >>> 𝟙)
bind f ≡⟨ cong bind (proj₂ .isIdentity)
bind f
forthRawEq : forthRaw (backRaw m) K.Monad.raw m forthRawEq : forthRaw (backRaw m) K.Monad.raw m
K.RawMonad.omap (forthRawEq _) = omap K.RawMonad.omap (forthRawEq _) = omap
@ -488,47 +489,44 @@ module _ {a b : Level} { : Category a b} where
fortheq m = K.Monad≡ (forthRawEq m) fortheq m = K.Monad≡ (forthRawEq m)
module _ (m : M.Monad) where module _ (m : M.Monad) where
open M.RawMonad (M.Monad.raw m) using (R ; Romap ; Rfmap ; pureNT ; joinNT) private
module KM = K.Monad (forth m) open M.Monad m
omapEq : KM.omap Romap module KM = K.Monad (forth m)
omapEq = refl module R = Functor R
omapEq : KM.omap Romap
omapEq = refl
D : (omap : Omap ) Romap omap Set _ bindEq : {X Y} {f : Arrow X (Romap Y)} KM.bind f bind f
D omap eq = (fmap' : Fmap omap) bindEq {X} {Y} {f} = begin
(λ i Fmap (eq i)) KM.bind f ≡⟨⟩
[ (λ f KM.fmap f) fmap' ] joinT Y Rfmap f ≡⟨⟩
bind f
-- The "base-case" for path induction on the family `D`. joinEq : {X} KM.join joinT X
d : D Romap λ _ Romap joinEq {X} = begin
d = res KM.join ≡⟨⟩
where KM.bind 𝟙 ≡⟨⟩
-- aka: bind 𝟙 ≡⟨⟩
res joinT X Rfmap 𝟙 ≡⟨ cong (λ φ _ φ) R.isIdentity
: (fmap : Fmap Romap) joinT X 𝟙 ≡⟨ proj₁ .isIdentity
(λ _ Fmap Romap) [ KM.fmap fmap ] joinT X
res fmap = begin
(λ f KM.fmap f) ≡⟨⟩
(λ f KM.bind (f >>> KM.pure)) ≡⟨ {!!}
(λ f fmap f)
-- This is sort of equivalent to `d` though the the order of fmapEq : {A B} KM.fmap {A} {B} Rfmap
-- quantification is different. `KM.fmap` is defined in terms of `Rfmap` fmapEq {A} {B} = funExt (λ f begin
-- (via `forth`) whereas in `d` above `fmap` is universally quantified. KM.fmap f ≡⟨⟩
-- KM.bind (f >>> KM.pure) ≡⟨⟩
-- I'm not sure `d` is provable. I believe `d'` should be, but I can also bind (f >>> pureT _) ≡⟨⟩
-- not prove it. Rfmap (f >>> pureT B) >>> joinT B ≡⟨⟩
d' : (λ i Fmap Romap) [ KM.fmap Rfmap ] Rfmap (f >>> pureT B) >>> joinT B ≡⟨ cong (λ φ φ >>> joinT B) R.isDistributive
d' = begin Rfmap f >>> Rfmap (pureT B) >>> joinT B ≡⟨ .isAssociative
(λ f KM.fmap f) ≡⟨⟩ joinT B Rfmap (pureT B) Rfmap f ≡⟨ cong (λ φ φ Rfmap f) (proj₂ isInverse)
(λ f KM.bind (f >>> KM.pure)) ≡⟨ {!!} 𝟙 Rfmap f ≡⟨ proj₂ .isIdentity
(λ f Rfmap f) Rfmap f
)
fmapEq : (λ i Fmap (omapEq i)) [ KM.fmap Rfmap ] rawEq : Functor.raw KM.R Functor.raw R
fmapEq = pathJ D d Romap refl Rfmap RawFunctor.func* (rawEq i) = omapEq i
RawFunctor.func→ (rawEq i) = fmapEq i
rawEq : Functor.raw KM.R Functor.raw R
RawFunctor.func* (rawEq i) = omapEq i
RawFunctor.func→ (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

View file

@ -23,6 +23,8 @@ module _ { ' : Level} ( : Category ') {A B obj : Object } whe
-- open IsProduct -- open IsProduct
-- TODO `isProp (Product ...)`
-- TODO `isProp (HasProducts ...)`
record Product { ' : Level} { : Category '} (A B : Object ) : Set ( ') where record Product { ' : Level} { : Category '} (A B : Object ) : Set ( ') where
no-eta-equality no-eta-equality
field field