mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-23 22:22:07 +03:00
Return error code from pclose
This commit is contained in:
parent
ac716c1dc7
commit
c1fc487bec
@ -11,7 +11,7 @@ prim__flush : FilePtr -> PrimIO Int
|
||||
%foreign support "idris2_popen"
|
||||
prim__popen : String -> String -> PrimIO FilePtr
|
||||
%foreign support "idris2_pclose"
|
||||
prim__pclose : FilePtr -> PrimIO ()
|
||||
prim__pclose : FilePtr -> PrimIO Int
|
||||
|
||||
||| Force a write of all user-space buffered data for the given `File`.
|
||||
|||
|
||||
@ -44,5 +44,5 @@ namespace Escaped
|
||||
|||
|
||||
||| @ fh the file handle to the stream to close/wait on
|
||||
export
|
||||
pclose : HasIO io => (fh : File) -> io ()
|
||||
pclose : HasIO io => (fh : File) -> io Int
|
||||
pclose (FHandle h) = primIO (prim__pclose h)
|
||||
|
@ -54,7 +54,7 @@ chezVersion chez = do
|
||||
| Left err => pure Nothing
|
||||
Right output <- fGetLine fh
|
||||
| Left err => pure Nothing
|
||||
pclose fh
|
||||
ignore $ pclose fh
|
||||
pure $ parseVersion output
|
||||
where
|
||||
cmd : String
|
||||
|
@ -456,9 +456,9 @@ hashFileWith (Just sha256sum) fileName
|
||||
| Left _ => err
|
||||
Right hashLine <- coreLift $ fGetLine fileHandle
|
||||
| Left _ =>
|
||||
do coreLift $ pclose fileHandle
|
||||
do ignore $ coreLift $ pclose fileHandle
|
||||
err
|
||||
coreLift $ pclose fileHandle
|
||||
ignore $ coreLift $ pclose fileHandle
|
||||
let w@(_::_) = words hashLine
|
||||
| Nil => err
|
||||
pure $ Just $ last w
|
||||
|
@ -107,13 +107,16 @@ void *idris2_popen(const char *cmd, const char *mode) {
|
||||
return f;
|
||||
}
|
||||
|
||||
void idris2_pclose(void *stream) {
|
||||
int idris2_pclose(void *stream) {
|
||||
#ifdef _WIN32
|
||||
int r = _pclose(stream);
|
||||
IDRIS2_VERIFY(r != -1, "pclose failed");
|
||||
return r;
|
||||
#else
|
||||
int r = pclose(stream);
|
||||
IDRIS2_VERIFY(WIFEXITED(r), "pclose failed");
|
||||
return WEXITSTATUS(r);
|
||||
#endif
|
||||
IDRIS2_VERIFY(r != -1, "pclose failed");
|
||||
}
|
||||
|
||||
// seek through the next newline, consuming and
|
||||
|
@ -23,7 +23,7 @@ int idris2_fileSize(FILE* h);
|
||||
int idris2_fpoll(FILE* f);
|
||||
|
||||
void *idris2_popen(const char *cmd, const char *mode);
|
||||
void idris2_pclose(void *stream);
|
||||
int idris2_pclose(void *stream);
|
||||
|
||||
// Seek through the next newline without
|
||||
// saving anything along the way
|
||||
|
@ -10,6 +10,8 @@ main = do
|
||||
Right contents <- fGetLine f
|
||||
| Left err => printLn err
|
||||
printLn $ trim contents
|
||||
0 <- pclose f
|
||||
| n => printLn n
|
||||
|
||||
let cmd : List String = if not isWindows
|
||||
then ["printf", "Hello, %s", "$PATH"]
|
||||
@ -19,3 +21,9 @@ main = do
|
||||
Right contents <- fGetLine f
|
||||
| Left err => printLn err
|
||||
printLn $ (ifThenElse isWindows trim id) contents
|
||||
0 <- pclose f
|
||||
| n => printLn n
|
||||
|
||||
Right f <- popen ["exit", "17"] Read
|
||||
| Left err => printLn err
|
||||
printLn !(pclose f)
|
||||
|
@ -1,4 +1,5 @@
|
||||
1/1: Building Popen (Popen.idr)
|
||||
Main> "Hello, world"
|
||||
"Hello, $PATH"
|
||||
17
|
||||
Main> Bye for now!
|
||||
|
@ -23,7 +23,7 @@ main = do
|
||||
putStrLn "opened"
|
||||
Right output <- fGetLine fh
|
||||
| Left err => printLn err
|
||||
pclose fh
|
||||
ignore $ pclose fh
|
||||
putStrLn "closed"
|
||||
let (idris2 ::: _) = split ((==) ',') output
|
||||
putStrLn idris2
|
||||
|
@ -22,7 +22,7 @@ main = do
|
||||
putStrLn "opened"
|
||||
Right output <- fGetLine fh
|
||||
| Left err => printLn err
|
||||
pclose fh
|
||||
ignore $ pclose fh
|
||||
putStrLn "closed"
|
||||
let [idris2, _] = split ((==) ',') output
|
||||
| _ => printLn "Unexpected result"
|
||||
|
Loading…
Reference in New Issue
Block a user