2020-05-18 15:59:07 +03:00
|
|
|
module System
|
|
|
|
|
Fix import loading
This was taking too long, and adding too many things, because it was
going too deep in the name of having everything accessible at the REPL
and for the compiler. So, it's done a bit differently now, only chasing
everything on a "full" load (i.e., final load at the REPL)
This has some effects:
+ As systems get bigger, load time gets better (on my machine, checking
Idris.Main now takes 52s from scratch, down from 76s)
+ You might find import errors that you didn't previously get, because
things were being imported that shouldn't have been. The new way is
correct!
An unfortunate effect is that sometimes you end up getting "undefined
name" errors even if you didn't explicitly use the name, because
sometimes a module uses a name from another module in a type, which then
gets exported, and eventually needs to be reduced. This mostly happens
because there is a compile time check that should be done which I
haven't implemented yet. That is, public export definitions should only
be allowed to use names that are also public export. I'll get to this
soon.
2020-05-27 17:49:03 +03:00
|
|
|
import public Data.So
|
2021-06-28 15:48:37 +03:00
|
|
|
import Data.String
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 18:38:32 +03:00
|
|
|
import public System.Escape
|
2021-11-03 19:07:33 +03:00
|
|
|
import System.File
|
2021-10-29 18:38:32 +03:00
|
|
|
|
2021-06-09 01:05:10 +03:00
|
|
|
%default total
|
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Shorthand for referring to the C support library
|
|
|
|
|||
|
|
|
|
||| @ fn the function name to refer to in the C support library
|
2022-01-19 09:43:03 +03:00
|
|
|
supportC : (fn : String) -> String
|
|
|
|
supportC fn = "C:\{fn}, libidris2_support, idris_support.h"
|
|
|
|
|
|
|
|
||| Shorthand for referring to the Node system support library
|
|
|
|
|||
|
|
|
|
||| @ fn the function name to refer to in the js/system_support.js file
|
|
|
|
supportNode : (fn : String) -> String
|
|
|
|
supportNode fn = "node:support:\{fn},support_system"
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Shorthand for referring to libc 6
|
|
|
|
|||
|
|
|
|
||| @ fn the function name to refer to in libc 6
|
|
|
|
libc : (fn : String) -> String
|
2020-05-18 15:59:07 +03:00
|
|
|
libc fn = "C:" ++ fn ++ ", libc 6"
|
|
|
|
|
2021-03-15 16:43:12 +03:00
|
|
|
-- `sleep` and `usleep` need to be tied to `blodwen-[u]sleep` for threading
|
|
|
|
-- reasons (see support/racket/support.rkt)
|
|
|
|
|
|
|
|
%foreign "scheme,racket:blodwen-sleep"
|
2022-01-19 09:43:03 +03:00
|
|
|
supportC "idris2_sleep"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__sleep : Int -> PrimIO ()
|
2021-03-15 16:43:12 +03:00
|
|
|
|
|
|
|
%foreign "scheme,racket:blodwen-usleep"
|
2022-01-19 09:43:03 +03:00
|
|
|
supportC "idris2_usleep"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__usleep : Int -> PrimIO ()
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Sleep for the specified number of seconds or, if signals are supported,
|
|
|
|
||| until an un-ignored signal arrives.
|
|
|
|
||| The exact wall-clock time slept might slighly differ depending on how busy
|
|
|
|
||| the system is and the resolution of the system's clock.
|
|
|
|
|||
|
|
|
|
||| @ sec the number of seconds to sleep for
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2021-10-29 19:58:29 +03:00
|
|
|
sleep : HasIO io => (sec : Int) -> io ()
|
2020-07-21 14:30:33 +03:00
|
|
|
sleep sec = primIO (prim__sleep sec)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Sleep for the specified number of microseconds or, if signals are supported,
|
|
|
|
||| until an un-ignored signal arrives.
|
|
|
|
||| The exact wall-clock time slept might slighly differ depending on how busy
|
|
|
|
||| the system is and the resolution of the system's clock.
|
|
|
|
|||
|
|
|
|
||| @ usec the number of microseconds to sleep for
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2021-10-29 19:58:29 +03:00
|
|
|
usleep : HasIO io => (usec : Int) -> So (usec >= 0) => io ()
|
|
|
|
usleep usec = primIO (prim__usleep usec)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-04-03 11:34:33 +03:00
|
|
|
-- Get the number of arguments
|
|
|
|
%foreign "scheme:blodwen-arg-count"
|
2022-01-19 09:43:03 +03:00
|
|
|
supportC "idris2_getArgCount"
|
2021-07-10 13:15:21 +03:00
|
|
|
"node:lambda:() => process.argv.length"
|
2021-04-03 11:34:33 +03:00
|
|
|
prim__getArgCount : PrimIO Int
|
|
|
|
|
|
|
|
-- Get argument number `n`
|
|
|
|
%foreign "scheme:blodwen-arg"
|
2022-01-19 09:43:03 +03:00
|
|
|
supportC "idris2_getArg"
|
2021-07-10 13:15:21 +03:00
|
|
|
"node:lambda:n => process.argv[n]"
|
2021-04-03 11:34:33 +03:00
|
|
|
prim__getArg : Int -> PrimIO String
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Retrieve the arguments to the program call, if there were any.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
getArgs : HasIO io => io (List String)
|
2021-04-03 11:34:33 +03:00
|
|
|
getArgs = do
|
|
|
|
n <- primIO prim__getArgCount
|
2021-04-07 18:33:59 +03:00
|
|
|
if n > 0
|
2021-06-16 17:22:30 +03:00
|
|
|
then for [0..n-1] $ primIO . prim__getArg
|
2021-04-03 11:34:33 +03:00
|
|
|
else pure []
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
%foreign libc "getenv"
|
2020-06-18 01:29:54 +03:00
|
|
|
"node:lambda: n => process.env[n]"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__getEnv : String -> PrimIO (Ptr String)
|
2020-06-18 01:29:54 +03:00
|
|
|
|
2022-01-19 09:43:03 +03:00
|
|
|
%foreign supportC "idris2_getEnvPair"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__getEnvPair : Int -> PrimIO (Ptr String)
|
2022-01-19 09:43:03 +03:00
|
|
|
%foreign supportC "idris2_setenv"
|
|
|
|
supportNode "setEnv"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__setEnv : String -> String -> Int -> PrimIO Int
|
2022-01-19 09:43:03 +03:00
|
|
|
%foreign supportC "idris2_unsetenv"
|
|
|
|
supportNode "unsetEnv"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__unsetEnv : String -> PrimIO Int
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Retrieve the specified environment variable's value string, or `Nothing` if
|
|
|
|
||| there is no such environment variable.
|
|
|
|
|||
|
|
|
|
||| @ var the name of the environment variable to look up
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2021-10-29 19:58:29 +03:00
|
|
|
getEnv : HasIO io => (var : String) -> io (Maybe String)
|
2020-05-18 15:59:07 +03:00
|
|
|
getEnv var
|
2020-07-21 14:30:33 +03:00
|
|
|
= do env <- primIO $ prim__getEnv var
|
2020-05-18 15:59:07 +03:00
|
|
|
if prim__nullPtr env /= 0
|
|
|
|
then pure Nothing
|
|
|
|
else pure (Just (prim__getString env))
|
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Retrieve all the key-value pairs of the environment variables, and return a
|
|
|
|
||| list containing them.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2021-06-09 01:05:10 +03:00
|
|
|
covering
|
Back to HasIO, remove MonadIO
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!)
2020-06-21 21:21:22 +03:00
|
|
|
getEnvironment : HasIO io => io (List (String, String))
|
2020-05-18 15:59:07 +03:00
|
|
|
getEnvironment = getAllPairs 0 []
|
|
|
|
where
|
|
|
|
splitEq : String -> (String, String)
|
|
|
|
splitEq str
|
|
|
|
= let (k, v) = break (== '=') str
|
|
|
|
(_, v') = break (/= '=') v in
|
|
|
|
(k, v')
|
|
|
|
|
2020-06-21 03:18:43 +03:00
|
|
|
getAllPairs : Int -> List String -> io (List (String, String))
|
2020-05-18 15:59:07 +03:00
|
|
|
getAllPairs n acc = do
|
2020-07-21 14:30:33 +03:00
|
|
|
envPair <- primIO $ prim__getEnvPair n
|
2020-05-18 15:59:07 +03:00
|
|
|
if prim__nullPtr envPair /= 0
|
|
|
|
then pure $ reverse $ map splitEq acc
|
|
|
|
else getAllPairs (n + 1) (prim__getString envPair :: acc)
|
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Add the specified variable with the given value string to the environment,
|
|
|
|
||| optionally overwriting any existing environment variable with the same name.
|
2022-01-19 09:43:03 +03:00
|
|
|
||| Returns True whether the value is set, overwritten, or not overwritten because
|
|
|
|
||| overwrite was False. Returns False if a system error occurred. You can `getErrno`
|
|
|
|
||| to check the error.
|
2021-10-29 19:58:29 +03:00
|
|
|
|||
|
|
|
|
||| @ var the name of the environment variable to set
|
|
|
|
||| @ val the value string to set the environment variable to
|
|
|
|
||| @ overwrite whether to overwrite the existing value if an environment
|
|
|
|
||| variable with the specified name already exists
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2021-10-29 19:58:29 +03:00
|
|
|
setEnv : HasIO io => (var : String) -> (val : String) -> (overwrite : Bool) ->
|
|
|
|
io Bool
|
2020-05-18 15:59:07 +03:00
|
|
|
setEnv var val overwrite
|
2020-07-21 14:30:33 +03:00
|
|
|
= do ok <- primIO $ prim__setEnv var val (if overwrite then 1 else 0)
|
2020-07-08 02:56:12 +03:00
|
|
|
pure $ ok == 0
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Delete the specified environment variable. Returns `True` either if the
|
|
|
|
||| value was deleted or if the value was not defined/didn't exist. Returns
|
2022-01-19 09:43:03 +03:00
|
|
|
||| `False` if a system error occurred. You can `getErrno` to check the error.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2021-10-29 19:58:29 +03:00
|
|
|
unsetEnv : HasIO io => (var : String) -> io Bool
|
2020-05-18 15:59:07 +03:00
|
|
|
unsetEnv var
|
2020-07-21 14:30:33 +03:00
|
|
|
= do ok <- primIO $ prim__unsetEnv var
|
2020-07-08 02:56:12 +03:00
|
|
|
pure $ ok == 0
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-06-27 09:27:22 +03:00
|
|
|
%foreign "C:idris2_system, libidris2_support, idris_system.h"
|
2022-01-19 09:43:03 +03:00
|
|
|
supportNode "spawnSync"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__system : String -> PrimIO Int
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Execute a shell command, returning its termination status or -1 if an error
|
|
|
|
||| occurred.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
system : HasIO io => String -> io Int
|
2020-07-21 14:30:33 +03:00
|
|
|
system cmd = primIO (prim__system cmd)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 18:38:32 +03:00
|
|
|
namespace Escaped
|
|
|
|
export
|
|
|
|
system : HasIO io => List String -> io Int
|
|
|
|
system = system . escapeCmd
|
|
|
|
|
2021-11-03 19:07:33 +03:00
|
|
|
||| Run a shell command, returning its stdout, and exit code.
|
|
|
|
export
|
|
|
|
covering
|
|
|
|
run : HasIO io => (cmd : String) -> io (String, Int)
|
|
|
|
run cmd = do
|
|
|
|
Right f <- popen cmd Read
|
|
|
|
| Left err => pure ("", 1)
|
|
|
|
Right resp <- fRead f
|
|
|
|
| Left err => pure ("", 1)
|
|
|
|
exitCode <- pclose f
|
|
|
|
pure (resp, exitCode)
|
|
|
|
|
|
|
|
namespace Escaped
|
|
|
|
export
|
|
|
|
covering
|
|
|
|
run : HasIO io => (cmd : List String) -> io (String, Int)
|
|
|
|
run = run . escapeCmd
|
|
|
|
|
2022-01-19 09:43:03 +03:00
|
|
|
%foreign supportC "idris2_time"
|
2021-08-29 23:43:55 +03:00
|
|
|
"javascript:lambda:() => Math.floor(new Date().getTime() / 1000)"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__time : PrimIO Int
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-06-29 01:09:55 +03:00
|
|
|
||| Return the number of seconds since epoch.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
Back to HasIO, remove MonadIO
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!)
2020-06-21 21:21:22 +03:00
|
|
|
time : HasIO io => io Integer
|
2020-07-21 14:30:33 +03:00
|
|
|
time = pure $ cast !(primIO prim__time)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2022-01-19 09:43:03 +03:00
|
|
|
%foreign supportC "idris2_getPID"
|
2021-12-22 17:25:21 +03:00
|
|
|
"node:lambda:() => process.pid"
|
2021-05-25 18:45:46 +03:00
|
|
|
prim__getPID : PrimIO Int
|
|
|
|
|
|
|
|
||| Get the ID of the currently running process.
|
|
|
|
export
|
|
|
|
getPID : HasIO io => io Int
|
|
|
|
getPID = primIO prim__getPID
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign libc "exit"
|
2021-07-10 13:15:21 +03:00
|
|
|
"node:lambda:c => process.exit(c)"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__exit : Int -> PrimIO ()
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| 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
|
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Exit the program normally, with the specified status.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
exitWith : HasIO io => ExitCode -> io a
|
2020-07-21 14:30:33 +03:00
|
|
|
exitWith ExitSuccess = primIO $ believe_me $ prim__exit 0
|
|
|
|
exitWith (ExitFailure code) = primIO $ believe_me $ prim__exit code
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-10-29 19:58:29 +03:00
|
|
|
||| Exit the program with status value 1, indicating failure.
|
|
|
|
||| If you want to specify a custom status value, see `exitWith`.
|
2020-05-18 15:59:07 +03:00
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
exitFailure : HasIO io => io a
|
2020-05-18 15:59:07 +03:00
|
|
|
exitFailure = exitWith (ExitFailure 1)
|
|
|
|
|
|
|
|
||| Exit the program after a successful run.
|
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
exitSuccess : HasIO io => io a
|
2020-05-18 15:59:07 +03:00
|
|
|
exitSuccess = exitWith ExitSuccess
|
2021-12-02 01:08:29 +03:00
|
|
|
|
|
|
|
||| Print the error message and call exitFailure
|
|
|
|
export
|
|
|
|
die : HasIO io => String -> io a
|
|
|
|
die str = do putStrLn str; exitFailure
|