mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-13 05:48:39 +03:00
Split HasIO into HasIO and MonadIO
For things which don't require (>>=), HasIO is fine, otherwise MonadIO gives access to the monad interface.
This commit is contained in:
parent
d12487f529
commit
28855088c2
@ -20,6 +20,9 @@ Language changes:
|
||||
|
||||
Library changes:
|
||||
|
||||
* `IO` operations in the `prelude` and `base` libraries now use the
|
||||
`HasIO` interface, rather than using `IO` directly.
|
||||
+ Also, added a `MonadIO` interface for operations which require `(>>=)`.
|
||||
* Experimental `Data.Linear.Array` added to `contrib`, support mutable linear
|
||||
arrays with constant time read/write, convertible to immutable arrays with
|
||||
constant time read.
|
||||
|
@ -14,19 +14,19 @@ max : IOArray elem -> Int
|
||||
max = maxSize
|
||||
|
||||
export
|
||||
newArray : Int -> IO (IOArray elem)
|
||||
newArray : MonadIO io => Int -> io (IOArray elem)
|
||||
newArray size
|
||||
= pure (MkIOArray size !(primIO (prim__newArray size Nothing)))
|
||||
|
||||
export
|
||||
writeArray : IOArray elem -> Int -> elem -> IO ()
|
||||
writeArray : MonadIO io => IOArray elem -> Int -> elem -> io ()
|
||||
writeArray arr pos el
|
||||
= if pos < 0 || pos >= max arr
|
||||
then pure ()
|
||||
else primIO (prim__arraySet (content arr) pos (Just el))
|
||||
|
||||
export
|
||||
readArray : IOArray elem -> Int -> IO (Maybe elem)
|
||||
readArray : MonadIO io => IOArray elem -> Int -> io (Maybe elem)
|
||||
readArray arr pos
|
||||
= if pos < 0 || pos >= max arr
|
||||
then pure Nothing
|
||||
@ -35,7 +35,8 @@ readArray arr pos
|
||||
-- Make a new array of the given size with the elements copied from the
|
||||
-- other array
|
||||
export
|
||||
newArrayCopy : (newsize : Int) -> IOArray elem -> IO (IOArray elem)
|
||||
newArrayCopy : MonadIO io =>
|
||||
(newsize : Int) -> IOArray elem -> io (IOArray elem)
|
||||
newArrayCopy newsize arr
|
||||
= do let newsize' = if newsize < max arr then max arr else newsize
|
||||
arr' <- newArray newsize'
|
||||
@ -44,7 +45,7 @@ newArrayCopy newsize arr
|
||||
where
|
||||
copyFrom : ArrayData (Maybe elem) ->
|
||||
ArrayData (Maybe elem) ->
|
||||
Int -> IO ()
|
||||
Int -> io ()
|
||||
copyFrom old new pos
|
||||
= if pos < 0
|
||||
then pure ()
|
||||
@ -53,10 +54,10 @@ newArrayCopy newsize arr
|
||||
assert_total (copyFrom old new (pos - 1))
|
||||
|
||||
export
|
||||
toList : IOArray elem -> IO (List (Maybe elem))
|
||||
toList : MonadIO io => IOArray elem -> io (List (Maybe elem))
|
||||
toList arr = iter 0 (max arr) []
|
||||
where
|
||||
iter : Int -> Int -> List (Maybe elem) -> IO (List (Maybe elem))
|
||||
iter : Int -> Int -> List (Maybe elem) -> io (List (Maybe elem))
|
||||
iter pos end acc
|
||||
= if pos >= end
|
||||
then pure (reverse acc)
|
||||
|
@ -13,7 +13,7 @@ data IORef : Type -> Type where
|
||||
MkRef : Mut a -> IORef a
|
||||
|
||||
export
|
||||
newIORef : HasIO io => a -> io (IORef a)
|
||||
newIORef : MonadIO io => a -> io (IORef a)
|
||||
newIORef val
|
||||
= do m <- primIO (prim__newIORef val)
|
||||
pure (MkRef m)
|
||||
@ -29,7 +29,7 @@ writeIORef : HasIO io => IORef a -> (1 val : a) -> io ()
|
||||
writeIORef (MkRef m) val = primIO (prim__writeIORef m val)
|
||||
|
||||
export
|
||||
modifyIORef : HasIO io => IORef a -> (a -> a) -> io ()
|
||||
modifyIORef : MonadIO io => IORef a -> (a -> a) -> io ()
|
||||
modifyIORef ref f
|
||||
= do val <- readIORef ref
|
||||
writeIORef ref (f val)
|
||||
|
@ -10,7 +10,7 @@ interface Ref m (r : Type -> Type) | m where
|
||||
writeRef : r a -> a -> m ()
|
||||
|
||||
export
|
||||
HasIO io => Ref io IORef where
|
||||
MonadIO io => Ref io IORef where
|
||||
newRef = newIORef
|
||||
readRef = readIORef
|
||||
writeRef = writeIORef
|
||||
|
@ -42,7 +42,7 @@ prim_setEnv : String -> String -> Int -> PrimIO Int
|
||||
prim_unsetEnv : String -> PrimIO Int
|
||||
|
||||
export
|
||||
getEnv : HasIO io => String -> io (Maybe String)
|
||||
getEnv : MonadIO io => String -> io (Maybe String)
|
||||
getEnv var
|
||||
= do env <- primIO $ prim_getEnv var
|
||||
if prim__nullPtr env /= 0
|
||||
@ -50,7 +50,7 @@ getEnv var
|
||||
else pure (Just (prim__getString env))
|
||||
|
||||
export
|
||||
getEnvironment : HasIO io => io (List (String, String))
|
||||
getEnvironment : MonadIO io => io (List (String, String))
|
||||
getEnvironment = getAllPairs 0 []
|
||||
where
|
||||
splitEq : String -> (String, String)
|
||||
@ -67,7 +67,7 @@ getEnvironment = getAllPairs 0 []
|
||||
else getAllPairs (n + 1) (prim__getString envPair :: acc)
|
||||
|
||||
export
|
||||
setEnv : HasIO io => String -> String -> Bool -> io Bool
|
||||
setEnv : MonadIO io => String -> String -> Bool -> io Bool
|
||||
setEnv var val overwrite
|
||||
= do ok <- primIO $ prim_setEnv var val (if overwrite then 1 else 0)
|
||||
if ok == 0
|
||||
@ -75,7 +75,7 @@ setEnv var val overwrite
|
||||
else pure False
|
||||
|
||||
export
|
||||
unsetEnv : HasIO io => String -> io Bool
|
||||
unsetEnv : MonadIO io => String -> io Bool
|
||||
unsetEnv var
|
||||
= do ok <- primIO $ prim_unsetEnv var
|
||||
if ok == 0
|
||||
@ -95,7 +95,7 @@ system cmd = primIO (prim_system cmd)
|
||||
prim_time : PrimIO Int
|
||||
|
||||
export
|
||||
time : HasIO io => io Integer
|
||||
time : MonadIO io => io Integer
|
||||
time = pure $ cast !(primIO prim_time)
|
||||
|
||||
%foreign libc "exit"
|
||||
|
@ -12,7 +12,7 @@ support fn = "C:" ++ fn ++ ", libidris2_support"
|
||||
%foreign support "idris2_fileErrno"
|
||||
prim_fileErrno : PrimIO Int
|
||||
|
||||
returnError : HasIO io => io (Either FileError a)
|
||||
returnError : MonadIO io => io (Either FileError a)
|
||||
returnError
|
||||
= do err <- primIO prim_fileErrno
|
||||
case err of
|
||||
@ -23,7 +23,7 @@ returnError
|
||||
4 => pure $ Left FileExists
|
||||
_ => pure $ Left (GenericFileError (err-5))
|
||||
|
||||
ok : HasIO io => a -> io (Either FileError a)
|
||||
ok : MonadIO io => a -> io (Either FileError a)
|
||||
ok x = pure (Right x)
|
||||
|
||||
%foreign support "idris2_currentDirectory"
|
||||
@ -52,7 +52,7 @@ data Directory : Type where
|
||||
MkDir : DirPtr -> Directory
|
||||
|
||||
export
|
||||
createDir : HasIO io => String -> io (Either FileError ())
|
||||
createDir : MonadIO io => String -> io (Either FileError ())
|
||||
createDir dir
|
||||
= do res <- primIO (prim_createDir dir)
|
||||
if res == 0
|
||||
@ -60,13 +60,13 @@ createDir dir
|
||||
else returnError
|
||||
|
||||
export
|
||||
changeDir : HasIO io => String -> io Bool
|
||||
changeDir : MonadIO io => String -> io Bool
|
||||
changeDir dir
|
||||
= do ok <- primIO (prim_changeDir dir)
|
||||
pure (ok == 0)
|
||||
|
||||
export
|
||||
currentDir : HasIO io => io (Maybe String)
|
||||
currentDir : MonadIO io => io (Maybe String)
|
||||
currentDir
|
||||
= do res <- primIO prim_currentDir
|
||||
if prim__nullPtr res /= 0
|
||||
@ -74,7 +74,7 @@ currentDir
|
||||
else pure (Just (prim__getString res))
|
||||
|
||||
export
|
||||
openDir : HasIO io => String -> io (Either FileError Directory)
|
||||
openDir : MonadIO io => String -> io (Either FileError Directory)
|
||||
openDir d
|
||||
= do res <- primIO (prim_openDir d)
|
||||
if prim__nullAnyPtr res /= 0
|
||||
@ -90,7 +90,7 @@ removeDir : HasIO io => String -> io ()
|
||||
removeDir dirName = primIO (prim_removeDir dirName)
|
||||
|
||||
export
|
||||
dirEntry : HasIO io => Directory -> io (Either FileError String)
|
||||
dirEntry : MonadIO io => Directory -> io (Either FileError String)
|
||||
dirEntry (MkDir d)
|
||||
= do res <- primIO (prim_dirEntry d)
|
||||
if prim__nullPtr res /= 0
|
||||
|
@ -84,7 +84,7 @@ data FileError = GenericFileError Int -- errno
|
||||
| PermissionDenied
|
||||
| FileExists
|
||||
|
||||
returnError : HasIO io => io (Either FileError a)
|
||||
returnError : MonadIO io => io (Either FileError a)
|
||||
returnError
|
||||
= do err <- primIO prim_fileErrno
|
||||
case err of
|
||||
@ -104,7 +104,7 @@ Show FileError where
|
||||
show PermissionDenied = "Permission Denied"
|
||||
show FileExists = "File Exists"
|
||||
|
||||
ok : HasIO io => a -> io (Either FileError a)
|
||||
ok : MonadIO io => a -> io (Either FileError a)
|
||||
ok x = pure (Right x)
|
||||
|
||||
public export
|
||||
@ -124,7 +124,7 @@ stderr : File
|
||||
stderr = FHandle prim__stderr
|
||||
|
||||
export
|
||||
openFile : HasIO io => String -> Mode -> io (Either FileError File)
|
||||
openFile : MonadIO io => String -> Mode -> io (Either FileError File)
|
||||
openFile f m
|
||||
= do res <- primIO (prim__open f (modeStr m) 0)
|
||||
if prim__nullAnyPtr res /= 0
|
||||
@ -136,13 +136,13 @@ closeFile : HasIO io => File -> io ()
|
||||
closeFile (FHandle f) = primIO (prim__close f)
|
||||
|
||||
export
|
||||
fileError : HasIO io => File -> io Bool
|
||||
fileError : MonadIO io => File -> io Bool
|
||||
fileError (FHandle f)
|
||||
= do x <- primIO $ prim_error f
|
||||
pure (x /= 0)
|
||||
|
||||
export
|
||||
fGetLine : HasIO io => (h : File) -> io (Either FileError String)
|
||||
fGetLine : MonadIO io => (h : File) -> io (Either FileError String)
|
||||
fGetLine (FHandle f)
|
||||
= do res <- primIO (prim__readLine f)
|
||||
if prim__nullPtr res /= 0
|
||||
@ -150,7 +150,7 @@ fGetLine (FHandle f)
|
||||
else ok (prim__getString res)
|
||||
|
||||
export
|
||||
fGetChars : HasIO io => (h : File) -> Int -> io (Either FileError String)
|
||||
fGetChars : MonadIO io => (h : File) -> Int -> io (Either FileError String)
|
||||
fGetChars (FHandle f) max
|
||||
= do res <- primIO (prim__readChars max f)
|
||||
if prim__nullPtr res /= 0
|
||||
@ -158,7 +158,7 @@ fGetChars (FHandle f) max
|
||||
else ok (prim__getString res)
|
||||
|
||||
export
|
||||
fGetChar : HasIO io => (h : File) -> io (Either FileError Char)
|
||||
fGetChar : MonadIO io => (h : File) -> io (Either FileError Char)
|
||||
fGetChar (FHandle h)
|
||||
= do c <- primIO (prim__readChar h)
|
||||
ferr <- primIO (prim_error h)
|
||||
@ -167,7 +167,7 @@ fGetChar (FHandle h)
|
||||
else ok (cast c)
|
||||
|
||||
export
|
||||
fPutStr : HasIO io => (h : File) -> String -> io (Either FileError ())
|
||||
fPutStr : MonadIO io => (h : File) -> String -> io (Either FileError ())
|
||||
fPutStr (FHandle f) str
|
||||
= do res <- primIO (prim__writeLine f str)
|
||||
if res == 0
|
||||
@ -175,23 +175,23 @@ fPutStr (FHandle f) str
|
||||
else ok ()
|
||||
|
||||
export
|
||||
fPutStrLn : HasIO io => (h : File) -> String -> io (Either FileError ())
|
||||
fPutStrLn : MonadIO io => (h : File) -> String -> io (Either FileError ())
|
||||
fPutStrLn f str = fPutStr f (str ++ "\n")
|
||||
|
||||
export
|
||||
fEOF : HasIO io => (h : File) -> io Bool
|
||||
fEOF : MonadIO io => (h : File) -> io Bool
|
||||
fEOF (FHandle f)
|
||||
= do res <- primIO (prim__eof f)
|
||||
pure (res /= 0)
|
||||
|
||||
export
|
||||
fflush : HasIO io => (h : File) -> io ()
|
||||
fflush : MonadIO io => (h : File) -> io ()
|
||||
fflush (FHandle f)
|
||||
= do primIO (prim__flush f)
|
||||
pure ()
|
||||
|
||||
export
|
||||
popen : HasIO io => String -> Mode -> io (Either FileError File)
|
||||
popen : MonadIO io => String -> Mode -> io (Either FileError File)
|
||||
popen f m = do
|
||||
ptr <- primIO (prim__popen f (modeStr m))
|
||||
if prim__nullAnyPtr ptr /= 0
|
||||
@ -203,7 +203,7 @@ pclose : HasIO io => File -> io ()
|
||||
pclose (FHandle h) = primIO (prim__pclose h)
|
||||
|
||||
export
|
||||
fileAccessTime : HasIO io => (h : File) -> io (Either FileError Int)
|
||||
fileAccessTime : MonadIO io => (h : File) -> io (Either FileError Int)
|
||||
fileAccessTime (FHandle f)
|
||||
= do res <- primIO (prim__fileAccessTime f)
|
||||
if res > 0
|
||||
@ -211,7 +211,7 @@ fileAccessTime (FHandle f)
|
||||
else returnError
|
||||
|
||||
export
|
||||
fileModifiedTime : HasIO io => (h : File) -> io (Either FileError Int)
|
||||
fileModifiedTime : MonadIO io => (h : File) -> io (Either FileError Int)
|
||||
fileModifiedTime (FHandle f)
|
||||
= do res <- primIO (prim__fileModifiedTime f)
|
||||
if res > 0
|
||||
@ -219,7 +219,7 @@ fileModifiedTime (FHandle f)
|
||||
else returnError
|
||||
|
||||
export
|
||||
fileStatusTime : HasIO io => (h : File) -> io (Either FileError Int)
|
||||
fileStatusTime : MonadIO io => (h : File) -> io (Either FileError Int)
|
||||
fileStatusTime (FHandle f)
|
||||
= do res <- primIO (prim__fileStatusTime f)
|
||||
if res > 0
|
||||
@ -227,7 +227,7 @@ fileStatusTime (FHandle f)
|
||||
else returnError
|
||||
|
||||
export
|
||||
removeFile : HasIO io => String -> io (Either FileError ())
|
||||
removeFile : MonadIO io => String -> io (Either FileError ())
|
||||
removeFile fname
|
||||
= do res <- primIO (prim__removeFile fname)
|
||||
if res == 0
|
||||
@ -235,7 +235,7 @@ removeFile fname
|
||||
else returnError
|
||||
|
||||
export
|
||||
fileSize : HasIO io => (h : File) -> io (Either FileError Int)
|
||||
fileSize : MonadIO io => (h : File) -> io (Either FileError Int)
|
||||
fileSize (FHandle f)
|
||||
= do res <- primIO (prim__fileSize f)
|
||||
if res >= 0
|
||||
@ -243,13 +243,13 @@ fileSize (FHandle f)
|
||||
else returnError
|
||||
|
||||
export
|
||||
fPoll : HasIO io => File -> io Bool
|
||||
fPoll : MonadIO io => File -> io Bool
|
||||
fPoll (FHandle f)
|
||||
= do p <- primIO (prim__fPoll f)
|
||||
pure (p > 0)
|
||||
|
||||
export
|
||||
readFile : HasIO io => String -> io (Either FileError String)
|
||||
readFile : MonadIO io => String -> io (Either FileError String)
|
||||
readFile file
|
||||
= do Right h <- openFile file Read
|
||||
| Left err => returnError
|
||||
@ -271,7 +271,7 @@ readFile file
|
||||
|
||||
||| Write a string to a file
|
||||
export
|
||||
writeFile : HasIO io =>
|
||||
writeFile : MonadIO io =>
|
||||
(filepath : String) -> (contents : String) ->
|
||||
io (Either FileError ())
|
||||
writeFile fn contents = do
|
||||
@ -307,7 +307,7 @@ mkMode p
|
||||
getMs = sum . map getM
|
||||
|
||||
export
|
||||
chmodRaw : HasIO io => String -> Int -> io (Either FileError ())
|
||||
chmodRaw : MonadIO io => String -> Int -> io (Either FileError ())
|
||||
chmodRaw fname p
|
||||
= do ok <- primIO $ prim__chmod fname p
|
||||
if ok == 0
|
||||
@ -315,5 +315,5 @@ chmodRaw fname p
|
||||
else returnError
|
||||
|
||||
export
|
||||
chmod : HasIO io => String -> Permissions -> io (Either FileError ())
|
||||
chmod : MonadIO io => String -> Permissions -> io (Either FileError ())
|
||||
chmod fname p = chmodRaw fname (mkMode p)
|
||||
|
@ -8,7 +8,7 @@ import System.File
|
||||
||| @ onInput the function to run on reading input, returning a String to
|
||||
||| output and a new state. Returns Nothing if the repl should exit
|
||||
export
|
||||
replWith : HasIO io =>
|
||||
replWith : MonadIO io =>
|
||||
(state : a) -> (prompt : String) ->
|
||||
(onInput : a -> String -> Maybe (String, a)) -> io ()
|
||||
replWith acc prompt fn
|
||||
@ -29,7 +29,7 @@ replWith acc prompt fn
|
||||
||| @ onInput the function to run on reading input, returning a String to
|
||||
||| output
|
||||
export
|
||||
repl : HasIO io =>
|
||||
repl : MonadIO io =>
|
||||
(prompt : String) -> (onInput : String -> String) -> io ()
|
||||
repl prompt fn
|
||||
= replWith () prompt (\x, s => Just (fn s, ()))
|
||||
|
@ -1564,9 +1564,15 @@ Monad IO where
|
||||
b >>= k = io_bind b k
|
||||
|
||||
public export
|
||||
interface Monad io => HasIO io where
|
||||
interface HasIO io where
|
||||
liftIO : (1 _ : IO a) -> io a
|
||||
|
||||
public export
|
||||
interface (Monad io, HasIO io) => MonadIO io where
|
||||
|
||||
public export %inline
|
||||
(Monad io, HasIO io) => MonadIO io where
|
||||
|
||||
public export %inline
|
||||
HasIO IO where
|
||||
liftIO x = x
|
||||
|
@ -1,6 +1,6 @@
|
||||
1/1: Building wcov (wcov.idr)
|
||||
Main> Main.tfoo is total
|
||||
Main> Main.wfoo is total
|
||||
Main> Main.wbar is not covering due to call to function Main.with block in 1387
|
||||
Main> Main.wbar is not covering due to call to function Main.with block in 1392
|
||||
Main> Main.wbar1 is total
|
||||
Main> Bye for now!
|
||||
|
@ -3,4 +3,4 @@ casetot.idr:12:1--13:1:main is not covering at:
|
||||
12 main : IO ()
|
||||
13 main = do
|
||||
|
||||
Calls non covering function Main.case block in 2114(619)
|
||||
Calls non covering function Main.case block in 2119(619)
|
||||
|
Loading…
Reference in New Issue
Block a user