mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-29 14:44:03 +03:00
Add local to MonadReader interface
This commit is contained in:
parent
6c3ab219bc
commit
f227735cec
@ -9,6 +9,9 @@ interface Monad m => MonadReader stateType (m : Type -> Type) | m where
|
||||
||| Get the context
|
||||
ask : m stateType
|
||||
|
||||
||| `local f c` runs the computation `c` in an environment modified by `f`.
|
||||
local : MonadReader stateType m => (stateType -> stateType) -> m a -> m a
|
||||
|
||||
||| The transformer on which the Reader monad is based
|
||||
public export
|
||||
record ReaderT (stateType : Type) (m: Type -> Type) (a: Type) where
|
||||
@ -44,6 +47,8 @@ public export
|
||||
implementation Monad m => MonadReader stateType (ReaderT stateType m) where
|
||||
ask = MkReaderT (\st => pure st)
|
||||
|
||||
local f (MkReaderT action) = MkReaderT (action . f)
|
||||
|
||||
public export
|
||||
implementation HasIO m => HasIO (ReaderT stateType m) where
|
||||
liftIO f = MkReaderT (\_ => liftIO f)
|
||||
|
Loading…
Reference in New Issue
Block a user