Rename zeta to pure
This commit is contained in:
parent
043641462d
commit
47882b1110
|
@ -74,12 +74,10 @@ module Kleisli {ℓa ℓb : Level} (ℂ : Category ℓa ℓb) where
|
||||||
field
|
field
|
||||||
RR : Object → Object
|
RR : Object → Object
|
||||||
-- Note name-change from [voe]
|
-- Note name-change from [voe]
|
||||||
ζ : {X : Object} → ℂ [ X , RR X ]
|
pure : {X : Object} → ℂ [ X , RR X ]
|
||||||
bind : {X Y : Object} → ℂ [ X , RR Y ] → ℂ [ RR X , RR Y ]
|
bind : {X Y : Object} → ℂ [ X , RR Y ] → ℂ [ RR X , RR Y ]
|
||||||
pure : {X : Object} → ℂ [ X , RR X ]
|
|
||||||
pure = ζ
|
|
||||||
fmap : ∀ {A B} → ℂ [ A , B ] → ℂ [ RR A , RR B ]
|
fmap : ∀ {A B} → ℂ [ A , B ] → ℂ [ RR A , RR B ]
|
||||||
fmap f = bind (ζ ∘ f)
|
fmap f = bind (pure ∘ f)
|
||||||
-- Why is (>>=) not implementable? - Because in e.g. the category of sets is
|
-- Why is (>>=) not implementable? - Because in e.g. the category of sets is
|
||||||
-- `m a` a set. This is not necessarily the case.
|
-- `m a` a set. This is not necessarily the case.
|
||||||
--
|
--
|
||||||
|
@ -126,12 +124,12 @@ module Kleisli {ℓa ℓb : Level} (ℂ : Category ℓa ℓb) where
|
||||||
bind ((fmap g ∘ pure) ∘ f) ≡⟨ cong bind (sym isAssociative) ⟩
|
bind ((fmap g ∘ pure) ∘ f) ≡⟨ cong bind (sym isAssociative) ⟩
|
||||||
bind
|
bind
|
||||||
(fmap g ∘ (pure ∘ f)) ≡⟨ sym lem ⟩
|
(fmap g ∘ (pure ∘ f)) ≡⟨ sym lem ⟩
|
||||||
bind (ζ ∘ g) ∘ bind (ζ ∘ f) ≡⟨⟩
|
bind (pure ∘ g) ∘ bind (pure ∘ f) ≡⟨⟩
|
||||||
fmap g ∘ fmap f ∎
|
fmap g ∘ fmap f ∎
|
||||||
where
|
where
|
||||||
open Category ℂ using (isAssociative)
|
open Category ℂ using (isAssociative)
|
||||||
lem : fmap g ∘ fmap f ≡ bind (fmap g ∘ (pure ∘ f))
|
lem : fmap g ∘ fmap f ≡ bind (fmap g ∘ (pure ∘ f))
|
||||||
lem = isDistributive (ζ ∘ g) (ζ ∘ f)
|
lem = isDistributive (pure ∘ g) (pure ∘ f)
|
||||||
|
|
||||||
record Monad : Set ℓ where
|
record Monad : Set ℓ where
|
||||||
field
|
field
|
||||||
|
@ -165,15 +163,15 @@ module _ {ℓa ℓb : Level} {ℂ : Category ℓa ℓb} where
|
||||||
RR : Object → Object
|
RR : Object → Object
|
||||||
RR = func* R
|
RR = func* R
|
||||||
|
|
||||||
ζ : {X : Object} → ℂ [ X , RR X ]
|
pure : {X : Object} → ℂ [ X , RR X ]
|
||||||
ζ {X} = η X
|
pure {X} = η X
|
||||||
|
|
||||||
bind : {X Y : Object} → ℂ [ X , RR Y ] → ℂ [ RR X , RR Y ]
|
bind : {X Y : Object} → ℂ [ X , RR Y ] → ℂ [ RR X , RR Y ]
|
||||||
bind {X} {Y} f = μ Y ∘ func→ R f
|
bind {X} {Y} f = μ Y ∘ func→ R f
|
||||||
|
|
||||||
forthRaw : K.RawMonad
|
forthRaw : K.RawMonad
|
||||||
Kraw.RR forthRaw = RR
|
Kraw.RR forthRaw = RR
|
||||||
Kraw.ζ forthRaw = ζ
|
Kraw.pure forthRaw = pure
|
||||||
Kraw.bind forthRaw = bind
|
Kraw.bind forthRaw = bind
|
||||||
|
|
||||||
module _ {raw : M.RawMonad} (m : M.IsMonad raw) where
|
module _ {raw : M.RawMonad} (m : M.IsMonad raw) where
|
||||||
|
@ -184,7 +182,7 @@ module _ {ℓa ℓb : Level} {ℂ : Category ℓa ℓb} where
|
||||||
|
|
||||||
isIdentity : IsIdentity
|
isIdentity : IsIdentity
|
||||||
isIdentity {X} = begin
|
isIdentity {X} = begin
|
||||||
bind ζ ≡⟨⟩
|
bind pure ≡⟨⟩
|
||||||
bind (η X) ≡⟨⟩
|
bind (η X) ≡⟨⟩
|
||||||
μ X ∘ func→ R (η X) ≡⟨ proj₂ isInverse ⟩
|
μ X ∘ func→ R (η X) ≡⟨ proj₂ isInverse ⟩
|
||||||
𝟙 ∎
|
𝟙 ∎
|
||||||
|
@ -192,7 +190,7 @@ module _ {ℓa ℓb : Level} {ℂ : Category ℓa ℓb} where
|
||||||
module R = Functor R
|
module R = Functor R
|
||||||
isNatural : IsNatural
|
isNatural : IsNatural
|
||||||
isNatural {X} {Y} f = begin
|
isNatural {X} {Y} f = begin
|
||||||
bind f ∘ ζ ≡⟨⟩
|
bind f ∘ pure ≡⟨⟩
|
||||||
bind f ∘ η X ≡⟨⟩
|
bind f ∘ η X ≡⟨⟩
|
||||||
μ Y ∘ R.func→ f ∘ η X ≡⟨ sym ℂ.isAssociative ⟩
|
μ Y ∘ R.func→ f ∘ η X ≡⟨ sym ℂ.isAssociative ⟩
|
||||||
μ Y ∘ (R.func→ f ∘ η X) ≡⟨ cong (λ φ → μ Y ∘ φ) (sym (ηN f)) ⟩
|
μ Y ∘ (R.func→ f ∘ η X) ≡⟨ cong (λ φ → μ Y ∘ φ) (sym (ηN f)) ⟩
|
||||||
|
@ -260,18 +258,18 @@ module _ {ℓa ℓb : Level} {ℂ : Category ℓa ℓb} where
|
||||||
|
|
||||||
rawR : RawFunctor ℂ ℂ
|
rawR : RawFunctor ℂ ℂ
|
||||||
RawFunctor.func* rawR = RR
|
RawFunctor.func* rawR = RR
|
||||||
RawFunctor.func→ rawR f = bind (ζ ∘ f)
|
RawFunctor.func→ rawR f = bind (pure ∘ f)
|
||||||
|
|
||||||
isFunctorR : IsFunctor ℂ ℂ rawR
|
isFunctorR : IsFunctor ℂ ℂ rawR
|
||||||
IsFunctor.isIdentity isFunctorR = begin
|
IsFunctor.isIdentity isFunctorR = begin
|
||||||
bind (ζ ∘ 𝟙) ≡⟨ cong bind (proj₁ ℂ.isIdentity) ⟩
|
bind (pure ∘ 𝟙) ≡⟨ cong bind (proj₁ ℂ.isIdentity) ⟩
|
||||||
bind ζ ≡⟨ isIdentity ⟩
|
bind pure ≡⟨ isIdentity ⟩
|
||||||
𝟙 ∎
|
𝟙 ∎
|
||||||
IsFunctor.isDistributive isFunctorR {f = f} {g} = begin
|
IsFunctor.isDistributive isFunctorR {f = f} {g} = begin
|
||||||
bind (ζ ∘ (g ∘ f)) ≡⟨⟩
|
bind (pure ∘ (g ∘ f)) ≡⟨⟩
|
||||||
fmap (g ∘ f) ≡⟨ fusion ⟩
|
fmap (g ∘ f) ≡⟨ fusion ⟩
|
||||||
fmap g ∘ fmap f ≡⟨⟩
|
fmap g ∘ fmap f ≡⟨⟩
|
||||||
bind (ζ ∘ g) ∘ bind (ζ ∘ f) ∎
|
bind (pure ∘ g) ∘ bind (pure ∘ f) ∎
|
||||||
|
|
||||||
R : Functor ℂ ℂ
|
R : Functor ℂ ℂ
|
||||||
Functor.raw R = rawR
|
Functor.raw R = rawR
|
||||||
|
@ -308,7 +306,7 @@ module _ {ℓa ℓb : Level} {ℂ : Category ℓa ℓb} where
|
||||||
open K.RawMonad (K.Monad.raw m)
|
open K.RawMonad (K.Monad.raw m)
|
||||||
forthRawEq : forthRaw (backRaw m) ≡ K.Monad.raw m
|
forthRawEq : forthRaw (backRaw m) ≡ K.Monad.raw m
|
||||||
K.RawMonad.RR (forthRawEq _) = RR
|
K.RawMonad.RR (forthRawEq _) = RR
|
||||||
K.RawMonad.ζ (forthRawEq _) = ζ
|
K.RawMonad.pure (forthRawEq _) = pure
|
||||||
-- stuck
|
-- stuck
|
||||||
K.RawMonad.bind (forthRawEq i) = {!!}
|
K.RawMonad.bind (forthRawEq i) = {!!}
|
||||||
|
|
||||||
|
|
Loading…
Reference in a new issue