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
|
2020-05-18 15:59:07 +03:00
|
|
|
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"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__sleep : Int -> PrimIO ()
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_usleep"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__usleep : Int -> PrimIO ()
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
sleep : HasIO io => Int -> io ()
|
2020-07-21 14:30:33 +03:00
|
|
|
sleep sec = primIO (prim__sleep sec)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
usleep : HasIO io => (x : Int) -> So (x >= 0) => io ()
|
2020-07-21 14:30:33 +03:00
|
|
|
usleep sec = primIO (prim__usleep sec)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
-- This one is going to vary for different back ends. Probably needs a
|
|
|
|
-- better convention. Will revisit...
|
|
|
|
%foreign "scheme:blodwen-args"
|
2020-06-18 01:29:54 +03:00
|
|
|
"node:lambda:() => __prim_js2idris_array(process.argv.slice(1))"
|
2020-05-18 15:59:07 +03:00
|
|
|
prim__getArgs : PrimIO (List String)
|
|
|
|
|
|
|
|
export
|
2020-06-21 03:18:43 +03:00
|
|
|
getArgs : HasIO io => io (List String)
|
2020-05-18 15:59:07 +03:00
|
|
|
getArgs = primIO prim__getArgs
|
|
|
|
|
|
|
|
%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
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
%foreign support "idris2_getEnvPair"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__getEnvPair : Int -> PrimIO (Ptr String)
|
2020-06-16 13:09:22 +03:00
|
|
|
%foreign support "idris2_setenv"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__setEnv : String -> String -> Int -> PrimIO Int
|
2020-06-16 13:09:22 +03:00
|
|
|
%foreign support "idris2_unsetenv"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__unsetEnv : String -> PrimIO Int
|
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
|
|
|
getEnv : HasIO io => 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))
|
|
|
|
|
|
|
|
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
|
|
|
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)
|
|
|
|
|
|
|
|
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
|
|
|
setEnv : HasIO io => String -> String -> 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
|
|
|
|
|
|
|
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
|
|
|
unsetEnv : HasIO io => 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
|
|
|
|
|
|
|
%foreign libc "system"
|
|
|
|
"scheme:blodwen-system"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__system : String -> PrimIO Int
|
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
|
|
|
|
|
|
|
%foreign support "idris2_time"
|
|
|
|
"scheme:blodwen-time"
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__time : PrimIO Int
|
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
|
|
|
|
|
|
|
%foreign libc "exit"
|
2020-06-18 01:29:54 +03:00
|
|
|
"node:lambda:c => process.exit(Number(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
|
|
|
|
|
|
|
|
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
|
|
|
|
|
|
|
||| Exit the program indicating failure.
|
|
|
|
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
|