Idris2/libs/base/Data/Ref.idr
Edwin Brady d12487f529 HasIO interface for IO actions
Also updates the Prelude and some base libraries to use HasIO rather
than using IO directly.
2020-06-21 01:18:43 +01:00

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