mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 11:05:17 +03:00
Replace prim__getArgs with prim__getArgsCount and prim__getArg
This spares us from having to build Idris data structures in the foreign code.
This commit is contained in:
parent
b2f08c0e81
commit
084997a880
@ -31,15 +31,23 @@ 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"
|
||||
"node:lambda:() => __prim_js2idris_array(process.argv.slice(1))"
|
||||
prim__getArgs : PrimIO (List String)
|
||||
-- Get the number of arguments
|
||||
%foreign "scheme:blodwen-arg-count"
|
||||
"node:lambda:() => process.argv.length"
|
||||
prim__getArgCount : PrimIO Int
|
||||
|
||||
-- Get argument number `n`
|
||||
%foreign "scheme:blodwen-arg"
|
||||
"node:lambda:n => process.argv[n]"
|
||||
prim__getArg : Int -> PrimIO String
|
||||
|
||||
export
|
||||
getArgs : HasIO io => io (List String)
|
||||
getArgs = primIO prim__getArgs
|
||||
getArgs = do
|
||||
n <- primIO prim__getArgCount
|
||||
if n > 0
|
||||
then for [0..n-1] (\x => primIO $ prim__getArg x)
|
||||
else pure []
|
||||
|
||||
%foreign libc "getenv"
|
||||
"node:lambda: n => process.env[n]"
|
||||
|
@ -358,12 +358,12 @@
|
||||
(define (blodwen-clock-second time) (time-second time))
|
||||
(define (blodwen-clock-nanosecond time) (time-nanosecond time))
|
||||
|
||||
(define (blodwen-args)
|
||||
(define (blodwen-build-args args)
|
||||
(if (null? args)
|
||||
(vector 0) ; Prelude.List
|
||||
(vector 1 (car args) (blodwen-build-args (cdr args)))))
|
||||
(blodwen-build-args (command-line)))
|
||||
|
||||
(define (blodwen-arg-count)
|
||||
(length (command-line)))
|
||||
|
||||
(define (blodwen-arg n)
|
||||
(if (< n (length (command-line))) (list-ref (command-line) n) ""))
|
||||
|
||||
(define (blodwen-hasenv var)
|
||||
(if (eq? (getenv var) #f) 0 1))
|
||||
|
@ -174,12 +174,12 @@
|
||||
(define (blodwen-time)
|
||||
(exact-floor (time->seconds (current-time))))
|
||||
|
||||
(define (blodwen-args)
|
||||
(define (blodwen-build-args args)
|
||||
(if (null? args)
|
||||
(vector 0) ; Prelude.List
|
||||
(vector 1 (car args) (blodwen-build-args (cdr args)))))
|
||||
(blodwen-build-args (cdr (command-line))))
|
||||
|
||||
(define (blodwen-arg-count)
|
||||
(length (command-line)))
|
||||
|
||||
(define (blodwen-arg n)
|
||||
(if (< n (length (command-line))) (list-ref (command-line) n) ""))
|
||||
|
||||
(define (blodwen-hasenv var)
|
||||
(if (getenv var #f) 1 0))
|
||||
|
@ -409,14 +409,13 @@
|
||||
(define (blodwen-clock-second time) (time-second time))
|
||||
(define (blodwen-clock-nanosecond time) (time-nanosecond time))
|
||||
|
||||
(define (blodwen-args)
|
||||
(define (blodwen-build-args args)
|
||||
(if (null? args)
|
||||
(vector 0) ; Prelude.List
|
||||
(vector 1 (car args) (blodwen-build-args (cdr args)))))
|
||||
(blodwen-build-args
|
||||
(cons (path->string (find-system-path 'run-file))
|
||||
(vector->list (current-command-line-arguments)))))
|
||||
|
||||
(define (blodwen-arg-count)
|
||||
(vector-length (current-command-line-arguments)))
|
||||
|
||||
(define (blodwen-arg n)
|
||||
(if (< n (vector-length (current-command-line-arguments)))
|
||||
(vector-ref (current-command-line-arguments) n) ""))
|
||||
|
||||
(define (blodwen-system cmd)
|
||||
(if (system cmd)
|
||||
|
Loading…
Reference in New Issue
Block a user