[ base ] Add modifyRef into the Ref interface

This commit is contained in:
Denis Buzdalov 2023-09-23 18:11:56 +03:00 committed by G. Allais
parent 8557b5b810
commit eddcbcdc75
2 changed files with 14 additions and 1 deletions

View File

@ -251,6 +251,9 @@
* Adds `Data.Vect.foldrImplGoLemma`.
* `Ref` interface from `Data.Ref` inherits `Monad` and was extended by a function
for value modification implemented through reading and writing by default.
#### System
* Changes `getNProcessors` to return the number of online processors rather than

View File

@ -7,11 +7,21 @@ import Control.Monad.State.Interface
%default total
public export
interface Ref m r | m where
interface Monad m => Ref m r | m where
newRef : {0 a : Type} -> a -> m (r a)
readRef : {0 a : Type} -> r a -> m a
writeRef : r a -> a -> m ()
||| Updates a value and returns the previous value
modifyRef : (a -> a) -> r a -> m a
modifyRef f ref = do
old <- readRef ref
writeRef ref (f old) $> old
public export
modifyRef_ : Ref m r => (a -> a) -> r a -> m ()
modifyRef_ = ignore .: modifyRef
export
HasIO io => Ref io IORef where
newRef = newIORef