Add RefC getArgs support

This commit is contained in:
Robert Wright 2021-05-17 14:07:53 +01:00 committed by G. Allais
parent f3aae06b28
commit cd3906645b
8 changed files with 43 additions and 3 deletions

View File

@ -31,11 +31,13 @@ usleep sec = primIO (prim__usleep sec)
-- Get the number of arguments
%foreign "scheme:blodwen-arg-count"
support "idris2_getArgCount"
"node:lambda:() => process.argv.length"
prim__getArgCount : PrimIO Int
-- Get argument number `n`
%foreign "scheme:blodwen-arg"
support "idris2_getArg"
"node:lambda:n => process.argv[n]"
prim__getArg : Int -> PrimIO String

View File

@ -928,12 +928,18 @@ header = do
fns <- get FunctionDefinitions
update OutfileText (appendL (initLines ++ headerLines ++ ["\n// function definitions"] ++ fns))
footer : {auto il : Ref IndentLevel Nat} -> {auto f : Ref OutfileText Output} -> Core ()
footer : {auto il : Ref IndentLevel Nat}
-> {auto f : Ref OutfileText Output}
-> {auto h : Ref HeaderFiles (SortedSet String)}
-> Core ()
footer = do
emit EmptyFC ""
emit EmptyFC " // main function"
emit EmptyFC "int main()"
emit EmptyFC "int main(int argc, char *argv[])"
emit EmptyFC "{"
if contains "idris_support.h" !(get HeaderFiles)
then emit EmptyFC " idris2_setArgs(argc, argv);"
else pure ()
emit EmptyFC " Value *mainExprVal = __mainExpression_0();"
emit EmptyFC " trampoline(mainExprVal);"
emit EmptyFC " return 0; // bye bye"

View File

@ -8,6 +8,9 @@
#include <stdlib.h>
#include <unistd.h>
int _argc;
char **_argv;
#ifdef _WIN32
extern char **_environ;
#include "windows/win_utils.h"
@ -80,6 +83,19 @@ int idris2_time() {
return time(NULL); // RTS needs to have 32 bit integers at least
}
void idris2_setArgs(int argc, char *argv[]) {
_argc = argc;
_argv = argv;
}
int idris2_getArgCount() {
return _argc;
}
char* idris2_getArg(int n) {
return _argv[n];
}
char* idris2_getEnvPair(int i) {
return *(environ + i);
}

View File

@ -17,6 +17,9 @@ void idris2_sleep(int sec);
void idris2_usleep(int usec);
int idris2_time();
int idris2_getArgCount();
void idris2_setArgs(int argc, char *argv[]);
char* idris2_getArg(int n);
char* idris2_getEnvPair(int i);
long idris2_getNProcessors();

View File

@ -224,7 +224,7 @@ refcTests : TestPool
refcTests = MkTestPool "Reference counting C backend" [C]
[ "refc001" , "refc002"
, "strings", "doubles"
, "buffer", "clock"
, "buffer", "clock", "args"
]
racketTests : TestPool

View File

@ -0,0 +1,6 @@
module TestArgs
import System
main : IO ()
main = getArgs >>= (putStrLn . show)

2
tests/refc/args/expected Normal file
View File

@ -0,0 +1,2 @@
["./build/exec/refc_args", "a", "b"]
["./build/exec/refc_args", "c"]

5
tests/refc/args/run Normal file
View File

@ -0,0 +1,5 @@
$1 --no-banner --no-color --console-width 0 --cg refc -o refc_args TestArgs.idr > /dev/null
./build/exec/refc_args a b
./build/exec/refc_args c
rm -rf build