mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-18 08:42:11 +03:00
d12487f529
Also updates the Prelude and some base libraries to use HasIO rather than using IO directly.
23 lines
416 B
Idris
23 lines
416 B
Idris
module Data.Ref
|
|
|
|
import public Data.IORef
|
|
import public Control.Monad.ST
|
|
|
|
public export
|
|
interface Ref m (r : Type -> Type) | m where
|
|
newRef : a -> m (r a)
|
|
readRef : r a -> m a
|
|
writeRef : r a -> a -> m ()
|
|
|
|
export
|
|
HasIO io => Ref io IORef where
|
|
newRef = newIORef
|
|
readRef = readIORef
|
|
writeRef = writeIORef
|
|
|
|
export
|
|
Ref (ST s) (STRef s) where
|
|
newRef = newSTRef
|
|
readRef = readSTRef
|
|
writeRef = writeSTRef
|