2021-01-21 15:24:40 +03:00
|
|
|
||| Note: The difference to a 'strict' RWST implementation is
|
2021-01-19 15:58:18 +03:00
|
|
|
||| that accumulation of values does not happen in the
|
2021-01-19 16:11:06 +03:00
|
|
|
||| Applicative and Monad instances but when invoking `Writer`-specific
|
2021-01-19 15:58:18 +03:00
|
|
|
||| functions like `writer` or `listen`.
|
2021-01-19 13:45:54 +03:00
|
|
|
module Control.Monad.RWS.CPS
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
import Control.Monad.Identity
|
|
|
|
import Control.Monad.Trans
|
|
|
|
|
2021-06-09 01:05:10 +03:00
|
|
|
%default total
|
|
|
|
|
2022-03-25 13:14:25 +03:00
|
|
|
||| A monad transformer extending an inner monad `m` with the ability to read
|
|
|
|
||| an environment of type `r`, collect an output of type `w` and update a
|
|
|
|
||| state of type `s`.
|
|
|
|
|||
|
|
|
|
||| This is equivalent to `ReaderT r (WriterT w (StateT s m)) a`, but fuses the
|
|
|
|
||| three layers.
|
2021-01-19 15:58:18 +03:00
|
|
|
public export
|
|
|
|
record RWST (r : Type) (w : Type) (s : Type) (m : Type -> Type) (a : Type) where
|
|
|
|
constructor MkRWST
|
|
|
|
unRWST : r -> s -> w -> m (a, s, w)
|
|
|
|
|
2022-03-25 13:14:25 +03:00
|
|
|
||| Unwrap an RWST computation as a function.
|
|
|
|
|||
|
|
|
|
||| This is the inverse of `rwsT`.
|
2021-01-19 15:58:18 +03:00
|
|
|
public export %inline
|
2022-06-14 16:06:37 +03:00
|
|
|
runRWST : Monoid w => r -> s -> RWST r w s m a -> m (a, s, w)
|
|
|
|
runRWST r s m = unRWST m r s neutral
|
2021-01-19 15:58:18 +03:00
|
|
|
|
2022-03-25 13:14:25 +03:00
|
|
|
||| Construct an RWST computation from a function.
|
|
|
|
|||
|
|
|
|
||| This is the inverse of `runRWST`.
|
2021-01-19 15:58:18 +03:00
|
|
|
public export %inline
|
2021-10-17 01:32:16 +03:00
|
|
|
rwsT : Semigroup w => Functor m => (r -> s -> m (a, s, w)) -> RWST r w s m a
|
2021-01-19 15:58:18 +03:00
|
|
|
rwsT f = MkRWST $ \r,s,w => (\(a,s',w') => (a,s',w <+> w')) <$> f r s
|
|
|
|
|
|
|
|
||| Evaluate a computation with the given initial state and environment,
|
|
|
|
||| returning the final value and output, discarding the final state.
|
|
|
|
public export %inline
|
2022-06-14 16:06:37 +03:00
|
|
|
evalRWST : Monoid w => Functor m => r -> s -> RWST r w s m a -> m (a,w)
|
|
|
|
evalRWST r s m = (\(a,_,w) => (a,w)) <$> runRWST r s m
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
||| Evaluate a computation with the given initial state and environment,
|
|
|
|
||| returning the final state and output, discarding the final value.
|
|
|
|
public export %inline
|
2022-06-14 16:06:37 +03:00
|
|
|
execRWST : Monoid w => Functor m => r -> s -> RWST r w s m a -> m (s,w)
|
|
|
|
execRWST r s m = (\(_,s',w) => (s',w)) <$> runRWST r s m
|
2021-01-19 15:58:18 +03:00
|
|
|
|
2022-03-25 13:14:25 +03:00
|
|
|
||| Map over the inner computation.
|
2021-01-19 15:58:18 +03:00
|
|
|
public export %inline
|
2021-01-19 16:11:06 +03:00
|
|
|
mapRWST : (Functor n, Monoid w, Semigroup w')
|
2021-01-19 15:58:18 +03:00
|
|
|
=> (m (a, s, w) -> n (b, s, w')) -> RWST r w s m a -> RWST r w' s n b
|
2021-08-10 21:24:32 +03:00
|
|
|
mapRWST f m = MkRWST $ \r,s,w =>
|
2022-06-14 16:06:37 +03:00
|
|
|
(\(a,s',w') => (a,s',w <+> w')) <$> f (runRWST r s m)
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
||| `withRWST f m` executes action `m` with an initial environment
|
|
|
|
||| and state modified by applying `f`.
|
|
|
|
public export %inline
|
|
|
|
withRWST : (r' -> s -> (r, s)) -> RWST r w s m a -> RWST r' w s m a
|
2021-08-10 21:24:32 +03:00
|
|
|
withRWST f m = MkRWST $ \r,s => uncurry (unRWST m) (f r s)
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
-- RWS Functions
|
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
|
|
|
|
||| A monad containing an environment of type `r`, output of type `w`
|
|
|
|
||| and an updatable state of type `s`.
|
2022-03-25 13:14:25 +03:00
|
|
|
|||
|
|
|
|
||| This is `RWST` applied to `Identity`.
|
2021-01-19 15:58:18 +03:00
|
|
|
public export
|
|
|
|
RWS : (r : Type) -> (w : Type) -> (s : Type) -> (a : Type) -> Type
|
|
|
|
RWS r w s = RWST r w s Identity
|
|
|
|
|
2022-03-25 13:14:25 +03:00
|
|
|
||| Unwrap an RWS computation as a function.
|
|
|
|
|||
|
|
|
|
||| This is the inverse of `rws`.
|
2021-01-19 15:58:18 +03:00
|
|
|
public export %inline
|
2022-06-14 16:06:37 +03:00
|
|
|
runRWS : Monoid w => r -> s -> RWS r w s a -> (a, s, w)
|
|
|
|
runRWS r s = runIdentity . runRWST r s
|
2021-01-19 15:58:18 +03:00
|
|
|
|
2022-03-25 13:14:25 +03:00
|
|
|
||| Construct an RWS computation from a function.
|
|
|
|
|||
|
|
|
|
||| This is the inverse of `runRWS`.
|
2021-01-19 15:58:18 +03:00
|
|
|
public export %inline
|
|
|
|
rws : Semigroup w => (r -> s -> (a, s, w)) -> RWS r w s a
|
2021-08-10 21:24:32 +03:00
|
|
|
rws f = MkRWST $ \r,s,w => let (a, s', w') = f r s
|
|
|
|
in Id (a, s', w <+> w')
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
||| Evaluate a computation with the given initial state and environment,
|
|
|
|
||| returning the final value and output, discarding the final state.
|
|
|
|
public export %inline
|
2022-06-14 16:06:37 +03:00
|
|
|
evalRWS : Monoid w => r -> s -> RWS r w s a -> (a,w)
|
|
|
|
evalRWS r s m = let (a,_,w) = runRWS r s m
|
2021-01-19 15:58:18 +03:00
|
|
|
in (a,w)
|
|
|
|
|
|
|
|
||| Evaluate a computation with the given initial state and environment,
|
|
|
|
||| returning the final state and output, discarding the final value.
|
|
|
|
public export %inline
|
2022-06-14 16:06:37 +03:00
|
|
|
execRWS : Monoid w => r -> s -> RWS r w s a -> (s,w)
|
|
|
|
execRWS r s m = let (_,s1,w) = runRWS r s m
|
2021-01-19 15:58:18 +03:00
|
|
|
in (s1,w)
|
|
|
|
|
|
|
|
||| Map the return value, final state and output of a computation using
|
|
|
|
||| the given function.
|
|
|
|
public export %inline
|
|
|
|
mapRWS : (Monoid w, Semigroup w')
|
|
|
|
=> ((a, s, w) -> (b, s, w')) -> RWS r w s a -> RWS r w' s b
|
2021-08-10 21:24:32 +03:00
|
|
|
mapRWS f = mapRWST $ \(Id p) => Id (f p)
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
||| `withRWS f m` executes action `m` with an initial environment
|
|
|
|
||| and state modified by applying `f`.
|
|
|
|
public export %inline
|
|
|
|
withRWS : (r' -> s -> (r, s)) -> RWS r w s a -> RWS r' w s a
|
|
|
|
withRWS = withRWST
|
|
|
|
|
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
-- Implementations
|
|
|
|
--------------------------------------------------------------------------------
|
|
|
|
|
|
|
|
public export %inline
|
|
|
|
Functor m => Functor (RWST r w s m) where
|
2021-08-10 21:24:32 +03:00
|
|
|
map f m = MkRWST $ \r,s,w => (\(a,s',w') => (f a,s',w')) <$> unRWST m r s w
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
public export %inline
|
|
|
|
Monad m => Applicative (RWST r w s m) where
|
2021-08-10 21:24:32 +03:00
|
|
|
pure a = MkRWST $ \_,s,w => pure (a,s,w)
|
2021-01-19 15:58:18 +03:00
|
|
|
MkRWST mf <*> MkRWST mx =
|
2021-08-10 21:24:32 +03:00
|
|
|
MkRWST $ \r,s,w => do (f,s1,w1) <- mf r s w
|
|
|
|
(a,s2,w2) <- mx r s1 w1
|
|
|
|
pure (f a,s2,w2)
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
public export %inline
|
|
|
|
(Monad m, Alternative m) => Alternative (RWST r w s m) where
|
2021-08-10 21:24:32 +03:00
|
|
|
empty = MkRWST $ \_,_,_ => empty
|
2021-08-10 14:44:51 +03:00
|
|
|
MkRWST m <|> mn = MkRWST $ \r,s,w => m r s w <|> unRWST mn r s w
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
public export %inline
|
|
|
|
Monad m => Monad (RWST r w s m) where
|
2021-08-10 21:24:32 +03:00
|
|
|
m >>= k = MkRWST $ \r,s,w => do (a,s1,w1) <- unRWST m r s w
|
|
|
|
unRWST (k a) r s1 w1
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
public export %inline
|
|
|
|
MonadTrans (RWST r w s) where
|
2021-08-10 21:24:32 +03:00
|
|
|
lift m = MkRWST $ \_,s,w => map (\a => (a,s,w)) m
|
2021-01-19 15:58:18 +03:00
|
|
|
|
|
|
|
public export %inline
|
|
|
|
HasIO m => HasIO (RWST r w s m) where
|
|
|
|
liftIO = lift . liftIO
|