mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-18 16:51:51 +03:00
dbdf7dab3d
Following a fairly detailed discussion on slack, the feeling is generally that it's better to have a single interface. While precision is nice, it doesn't appear to buy us anything here. If that turns out to be wrong, or limiting somehow, we can revisit it later. Also: - it's easier for backend authors if the type of IO operations is slightly less restrictive. For example, if it's in HasIO, that limits alternative implementations, which might be awkward for some alternative back ends. - it's one less extra detail to learn. This is minor, but there needs to be a clear advantage if there's more detail to learn. - It is difficult to think of an underlying type that can't have a Monad instance (I have personally never encountered one - if they turns out to exist, again, we can revisit!)
130 lines
3.4 KiB
Idris
130 lines
3.4 KiB
Idris
module System
|
|
|
|
import public Data.So
|
|
import Data.List
|
|
import Data.Strings
|
|
|
|
support : String -> String
|
|
support fn = "C:" ++ fn ++ ", libidris2_support"
|
|
|
|
libc : String -> String
|
|
libc fn = "C:" ++ fn ++ ", libc 6"
|
|
|
|
%foreign support "idris2_sleep"
|
|
prim_sleep : Int -> PrimIO ()
|
|
%foreign support "idris2_usleep"
|
|
prim_usleep : Int -> PrimIO ()
|
|
|
|
export
|
|
sleep : HasIO io => Int -> io ()
|
|
sleep sec = primIO (prim_sleep sec)
|
|
|
|
export
|
|
usleep : HasIO io => (x : Int) -> So (x >= 0) => io ()
|
|
usleep sec = primIO (prim_usleep sec)
|
|
|
|
-- This one is going to vary for different back ends. Probably needs a
|
|
-- better convention. Will revisit...
|
|
%foreign "scheme:blodwen-args"
|
|
prim__getArgs : PrimIO (List String)
|
|
|
|
export
|
|
getArgs : HasIO io => io (List String)
|
|
getArgs = primIO prim__getArgs
|
|
|
|
%foreign libc "getenv"
|
|
prim_getEnv : String -> PrimIO (Ptr String)
|
|
%foreign support "idris2_getEnvPair"
|
|
prim_getEnvPair : Int -> PrimIO (Ptr String)
|
|
%foreign support "idris2_setenv"
|
|
prim_setEnv : String -> String -> Int -> PrimIO Int
|
|
%foreign support "idris2_unsetenv"
|
|
prim_unsetEnv : String -> PrimIO Int
|
|
|
|
export
|
|
getEnv : HasIO io => String -> io (Maybe String)
|
|
getEnv var
|
|
= do env <- primIO $ prim_getEnv var
|
|
if prim__nullPtr env /= 0
|
|
then pure Nothing
|
|
else pure (Just (prim__getString env))
|
|
|
|
export
|
|
getEnvironment : HasIO io => io (List (String, String))
|
|
getEnvironment = getAllPairs 0 []
|
|
where
|
|
splitEq : String -> (String, String)
|
|
splitEq str
|
|
= let (k, v) = break (== '=') str
|
|
(_, v') = break (/= '=') v in
|
|
(k, v')
|
|
|
|
getAllPairs : Int -> List String -> io (List (String, String))
|
|
getAllPairs n acc = do
|
|
envPair <- primIO $ prim_getEnvPair n
|
|
if prim__nullPtr envPair /= 0
|
|
then pure $ reverse $ map splitEq acc
|
|
else getAllPairs (n + 1) (prim__getString envPair :: acc)
|
|
|
|
export
|
|
setEnv : HasIO 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
|
|
then pure True
|
|
else pure False
|
|
|
|
export
|
|
unsetEnv : HasIO io => String -> io Bool
|
|
unsetEnv var
|
|
= do ok <- primIO $ prim_unsetEnv var
|
|
if ok == 0
|
|
then pure True
|
|
else pure False
|
|
|
|
%foreign libc "system"
|
|
"scheme:blodwen-system"
|
|
prim_system : String -> PrimIO Int
|
|
|
|
export
|
|
system : HasIO io => String -> io Int
|
|
system cmd = primIO (prim_system cmd)
|
|
|
|
%foreign support "idris2_time"
|
|
"scheme:blodwen-time"
|
|
prim_time : PrimIO Int
|
|
|
|
export
|
|
time : HasIO io => io Integer
|
|
time = pure $ cast !(primIO prim_time)
|
|
|
|
%foreign libc "exit"
|
|
prim_exit : Int -> PrimIO ()
|
|
|
|
||| Programs can either terminate successfully, or end in a caught
|
|
||| failure.
|
|
public export
|
|
data ExitCode : Type where
|
|
||| Terminate successfully.
|
|
ExitSuccess : ExitCode
|
|
||| Program terminated for some prescribed reason.
|
|
|||
|
|
||| @errNo A non-zero numerical value indicating failure.
|
|
||| @prf Proof that the int value is non-zero.
|
|
ExitFailure : (errNo : Int) -> (So (not $ errNo == 0)) => ExitCode
|
|
|
|
export
|
|
exitWith : HasIO io => ExitCode -> io a
|
|
exitWith ExitSuccess = primIO $ believe_me $ prim_exit 0
|
|
exitWith (ExitFailure code) = primIO $ believe_me $ prim_exit code
|
|
|
|
||| Exit the program indicating failure.
|
|
export
|
|
exitFailure : HasIO io => io a
|
|
exitFailure = exitWith (ExitFailure 1)
|
|
|
|
||| Exit the program after a successful run.
|
|
export
|
|
exitSuccess : HasIO io => io a
|
|
exitSuccess = exitWith ExitSuccess
|