{-# OPTIONS --cubical-compatible --safe #-}
open import Level
open import Algebra
module Data.These.Effectful.Right (a : Level) {c ℓ} (W : Semigroup c ℓ) where
open Semigroup W
open import Data.These.Effectful.Right.Base a Carrier public
open import Data.These.Base
open import Effect.Applicative
open import Effect.Monad
module _ {a b} {A : Set a} {B : Set b} where
applicative : RawApplicative Theseᵣ
applicative = record
{ rawFunctor = functor
; pure = this
; _<*>_ = ap
} where
ap : ∀ {A B}→ Theseᵣ (A → B) → Theseᵣ A → Theseᵣ B
ap (this f) t = map₁ f t
ap (that w) t = that w
ap (these f w) t = map f (w ∙_) t
monad : RawMonad Theseᵣ
monad = record
{ rawApplicative = applicative
; _>>=_ = bind
} where
bind : ∀ {A B} → Theseᵣ A → (A → Theseᵣ B) → Theseᵣ B
bind (this t) f = f t
bind (that w) f = that w
bind (these t w) f = map₂ (w ∙_) (f t)