Idris2/tests/node/node020/Popen.idr
Mathew Polzin ad12f8335c
[ new ] popen/pclose for the NodeJS backend. (#2857)
* implement popen and pclose (to an extent) for NodeJS

* bring node020 back into tests.

* ah, I see what was being done here. Fix the idris for the test.

* fix test's unreachable clause warning

* fix expectation

* Add note to CHANGELOG

* small tweaks to get popen into merge-ready state.
2023-01-30 09:38:42 -06:00

30 lines
832 B
Idris

import System
import System.File
import System.Info
import Data.String
import Data.List1
windowsPath : String -> String
windowsPath path =
let replaceSlashes : List Char -> List Char
replaceSlashes [] = []
replaceSlashes ('/' :: cs) = '\\' :: replaceSlashes cs
replaceSlashes (c :: cs) = c :: replaceSlashes cs
in
pack $ replaceSlashes (unpack path)
main : IO ()
main = do
Just cmd <- getEnv "POPEN_CMD"
| Nothing => putStrLn "POPEN_CMD env var not set"
let cmd = if isWindows then windowsPath cmd else cmd
Right fh <- popen cmd Read
| Left err => printLn err
putStrLn "opened"
Right output <- fGetLine fh
| Left err => printLn err
ignore $ pclose fh
putStrLn "closed"
let (idris2 ::: _) = split ((==) ',') output
putStrLn idris2