mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 02:23:44 +03:00
Add System run function
This commit is contained in:
parent
dc47df688c
commit
4732486bbc
@ -5,7 +5,7 @@ import Data.List
|
||||
import Data.String
|
||||
|
||||
import public System.Escape
|
||||
import System.File.Process
|
||||
import System.File
|
||||
|
||||
%default total
|
||||
|
||||
@ -153,6 +153,24 @@ namespace Escaped
|
||||
system : HasIO io => List String -> io Int
|
||||
system = system . escapeCmd
|
||||
|
||||
||| 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
|
||||
|
||||
%foreign support "idris2_time"
|
||||
"javascript:lambda:() => Math.floor(new Date().getTime() / 1000)"
|
||||
prim__time : PrimIO Int
|
||||
|
19
tests/base/system_run/Run.idr
Normal file
19
tests/base/system_run/Run.idr
Normal file
@ -0,0 +1,19 @@
|
||||
import Data.String
|
||||
|
||||
import System
|
||||
import System.Info
|
||||
|
||||
main : IO ()
|
||||
main = do
|
||||
(contents, 0) <- run $ "echo " ++ escapeArg "Hello, world"
|
||||
| (_, err) => printLn err
|
||||
printLn $ trim contents
|
||||
|
||||
let cmd : List String = if not isWindows
|
||||
then ["printf", "Hello, %s", "$PATH"]
|
||||
else ["echo", "Hello, $PATH"]
|
||||
(contents, 0) <- run cmd
|
||||
| (_, err) => printLn err
|
||||
printLn $ trim contents
|
||||
|
||||
printLn !(run ["exit", "17"])
|
5
tests/base/system_run/expected
Normal file
5
tests/base/system_run/expected
Normal file
@ -0,0 +1,5 @@
|
||||
1/1: Building Run (Run.idr)
|
||||
Main> "Hello, world"
|
||||
"Hello, $PATH"
|
||||
("", 17)
|
||||
Main> Bye for now!
|
2
tests/base/system_run/input
Normal file
2
tests/base/system_run/input
Normal file
@ -0,0 +1,2 @@
|
||||
:exec main
|
||||
:q
|
3
tests/base/system_run/run
Normal file
3
tests/base/system_run/run
Normal file
@ -0,0 +1,3 @@
|
||||
rm -rf build
|
||||
|
||||
$1 --no-color --console-width 0 --no-banner Run.idr < input
|
Loading…
Reference in New Issue
Block a user