removed MonadReader functions from RWS.CPS.idr

This commit is contained in:
stefan-hoeck 2021-01-20 05:45:32 +01:00
parent 2776eaa1b1
commit 6f2e358f1b
2 changed files with 4 additions and 23 deletions

View File

@ -132,25 +132,6 @@ public export %inline
HasIO m => HasIO (RWST r w s m) where
liftIO = lift . liftIO
--------------------------------------------------------------------------------
-- Reader Operations
--------------------------------------------------------------------------------
||| Retrieve a function of the current environment.
public export %inline
asks : Applicative m => (r -> a) -> RWST r w s m a
asks f = MkRWST \r,s,w => pure (f r,s,w)
||| Fetch the value of the environment.
public export %inline
ask : Applicative m => RWST r w s m r
ask = asks id
||| Execute a computation in a modified environment
public export %inline
local : (r -> r) -> RWST r w s m a -> RWST r w s m a
local f m = MkRWST \r,s,w => unRWST m (f r) s w
--------------------------------------------------------------------------------
-- Writer Operations
--------------------------------------------------------------------------------

View File

@ -4,7 +4,7 @@ import Control.Monad.Maybe
import Control.Monad.Error.Either
import Control.Monad.Reader.Reader
import Control.Monad.State.State
import Control.Monad.RWS.CPS as RWS
import Control.Monad.RWS.CPS
import Control.Monad.Trans
import Control.Monad.Writer.CPS
@ -21,7 +21,7 @@ interface Monad m => MonadReader stateType m | m where
||| Evaluate a function in the context held by this computation
public export
asks : MonadReader stateType m => (stateType -> a) -> m a
asks f = ask >>= pure . f
asks f = map f ask
--------------------------------------------------------------------------------
-- Implementations
@ -35,8 +35,8 @@ Monad m => MonadReader stateType (ReaderT stateType m) where
public export %inline
Monad m => MonadReader r (RWST r w s m) where
ask = RWS.ask
local = RWS.local
ask = MkRWST \r,s,w => pure (r,s,w)
local f m = MkRWST \r,s,w => unRWST m (f r) s w
public export %inline
MonadReader r m => MonadReader r (EitherT e m) where