remove MonadWriter functions from Writer.CPS module

This commit is contained in:
stefan-hoeck 2021-01-20 05:41:01 +01:00
parent b6e5ba0830
commit 2776eaa1b1
2 changed files with 15 additions and 53 deletions

View File

@ -105,48 +105,3 @@ MonadTrans (WriterT w) where
public export %inline
HasIO m => HasIO (WriterT w m) where
liftIO = lift . liftIO
--------------------------------------------------------------------------------
-- Writer Operations
--------------------------------------------------------------------------------
||| Construct a writer computation from a (result, output) pair.
public export %inline
writer : (Semigroup w, Applicative m) => (a, w) -> WriterT w m a
writer = writerT . pure
||| `tell w` is an action that produces the output `w`.
public export %inline
tell : (Semigroup w, Applicative m) => w -> WriterT w m ()
tell w' = writer ((), w')
||| `listens f m` is an action that executes the action `m` and adds
||| the result of applying `f` to the output to the value of the computation.
public export %inline
listens : (Monoid w, Functor m)
=> (w -> b) -> WriterT w m a -> WriterT w m (a, b)
listens f m = MkWriterT \w =>
(\(a,w') => ((a,f w'),w <+> w')) <$> runWriterT m
||| `listen m` is an action that executes the action `m` and adds its
||| output to the value of the computation.
public export %inline
listen : (Monoid w, Functor m) => WriterT w m a -> WriterT w m (a, w)
listen = listens id
||| `pass m` is an action that executes the action `m`, which returns
||| a value and a function, and returns the value, applying the function
||| to the output.
public export %inline
pass : (Monoid w, Semigroup w', Functor m)
=> WriterT w m (a, w -> w') -> WriterT w' m a
pass m = MkWriterT \w =>
(\((a,f),w') => (a,w <+> f w')) <$> runWriterT m
||| `censor f m` is an action that executes the action `m` and
||| applies the function `f` to its output, leaving the return value
||| unchanged.
public export %inline
censor : (Monoid w, Functor m) => (w -> w) -> WriterT w m a -> WriterT w m a
censor f m = MkWriterT \w =>
(\(a,w') => (a,w <+> f w')) <$> runWriterT m

View File

@ -41,23 +41,30 @@ interface (Monoid w, Monad m) => MonadWriter w m | m where
||| the result of applying @f@ to the output to the value of the computation.
public export
listens : MonadWriter w m => (w -> b) -> m a -> m (a, b)
listens f m = do (a, w) <- listen m
pure (a, f w)
listens f = map (\(a,w) => (a,f w)) . listen
||| `censor f m` is an action that executes the action `m` and
||| applies the function `f` to its output, leaving the return value
||| unchanged.
public export
censor : MonadWriter w m => (w -> w) -> m a -> m a
censor f m = pass $ do a <- m
pure (a, f)
censor f = pass . map (\a => (a,f))
--------------------------------------------------------------------------------
-- Implementations
--------------------------------------------------------------------------------
public export %inline
(Monoid w, Monad m) => MonadWriter w (WriterT w m) where
writer = Writer.writer
tell = Writer.tell
listen = Writer.listen
pass = Writer.pass
writer = writerT . pure
listen m = MkWriterT \w =>
(\(a,w') => ((a,w'),w <+> w')) <$> runWriterT m
tell w' = writer ((), w')
pass m = MkWriterT \w =>
(\((a,f),w') => (a,w <+> f w')) <$> runWriterT m
public export %inline
(Monoid w, Monad m) => MonadWriter w (RWST r w s m) where