Skip to content

Commit

Permalink
Merge pull request #3170 from buzden/make-able-wait-popen2
Browse files Browse the repository at this point in the history
[ re #2944, fix ] Make `popen2` be able to not to leak system resources
  • Loading branch information
andrevidela committed Dec 20, 2023
2 parents 4cb45da + b495fe8 commit 3502f4a
Show file tree
Hide file tree
Showing 7 changed files with 147 additions and 6 deletions.
4 changes: 4 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -254,6 +254,10 @@
* `Ref` interface from `Data.Ref` inherits `Monad` and was extended by a function
for value modification implemented through reading and writing by default.

* A function `popen2Wait` was added to wait for the process started with `popen2`
function and clean up all system resources (to not to leave zombie processes in
particular).

#### System

* Changes `getNProcessors` to return the number of online processors rather than
Expand Down
34 changes: 32 additions & 2 deletions libs/base/System/File/Process.idr
Original file line number Diff line number Diff line change
Expand Up @@ -16,14 +16,24 @@ prim__popen : String -> String -> PrimIO FilePtr
supportNode "pclose"
prim__pclose : FilePtr -> PrimIO Int

data Popen2Result : Type where
data Popen2Result : Type where [external]

%foreign supportC "idris2_popen2"
prim__popen2 : String -> PrimIO (Ptr Popen2Result)

%foreign supportC "idris2_popen2WaitByPid"
covering -- significantly non-total
prim__popen2WaitByPid : Int -> PrimIO Int
%foreign supportC "idris2_popen2WaitByHandler"
covering -- significantly non-total
prim__popen2WaitByHandler : AnyPtr -> PrimIO Int

%foreign supportC "idris2_popen2ChildPid"
prim__popen2ChildPid : Ptr Popen2Result -> PrimIO Int

%foreign supportC "idris2_popen2ChildHandler"
prim__popen2ChildHandler : Ptr Popen2Result -> PrimIO AnyPtr

%foreign supportC "idris2_popen2FileIn"
prim__popen2FileIn : Ptr Popen2Result -> PrimIO FilePtr

Expand Down Expand Up @@ -71,6 +81,8 @@ record SubProcess where
constructor MkSubProcess
||| Process id of the spawned process
pid : Int
||| The way to manipulate the spawned process, for systems where pid is not enough for this
handler : AnyPtr
||| The input stream of the spawned process
input : File
||| The output stream of the spawned process
Expand All @@ -80,6 +92,8 @@ record SubProcess where
||| given command-string using the '-c' flag, in a new process. On success
||| a SubProcess is returned which holds the process id and file handles
||| for input and output.
||| You should call `popen2Wait` after you've done with the child process
||| in order to clean up all system resources.
|||
||| IMPORTANT: You may deadlock if write to a process which is waiting to flush
||| its output. It is recommended to read and write from separate threads.
Expand All @@ -93,10 +107,26 @@ popen2 cmd = do
then returnError
else do
pid <- primIO (prim__popen2ChildPid ptr)
handle <- primIO (prim__popen2ChildHandler ptr)
input <- primIO (prim__popen2FileIn ptr)
output <- primIO (prim__popen2FileOut ptr)
free (prim__forgetPtr ptr)
pure $ Right (MkSubProcess pid (FHandle input) (FHandle output))
pure $ Right (MkSubProcess pid handle (FHandle input) (FHandle output))

||| Blocks and waits until the process created by `popen2` finished
|||
||| This function relates to `popen2` like `pclose` relates to `popen`.
||| Returns exit code of the process being waited.
||| IMPORTANT: this function mustn't be called twice for the same argument.
|||
||| Support of this function in the backends must be the same as for `popen2`.
export
covering -- significantly non-total
popen2Wait : HasIO io => SubProcess -> io Int
popen2Wait sp = primIO $
if prim__nullAnyPtr sp.handler /= 0
then prim__popen2WaitByPid sp.pid
else prim__popen2WaitByHandler sp.handler

namespace Escaped
export
Expand Down
50 changes: 49 additions & 1 deletion support/c/idris_file.c
Original file line number Diff line number Diff line change
Expand Up @@ -275,6 +275,9 @@ FILE *idris2_stdout() { return stdout; }
FILE *idris2_stderr() { return stderr; }

struct child_process {
#ifdef _WIN32
HANDLE hProcess;
#endif
pid_t pid;
FILE *in;
FILE *out;
Expand Down Expand Up @@ -320,10 +323,12 @@ struct child_process *idris2_popen2(char *cmd) {
int out_fd = _open_osfhandle((intptr_t)pipes[0], _O_RDONLY);
CloseHandle(pipes[1]);
CloseHandle(pipes[2]);
CloseHandle(pi.hProcess);
// We close thread handle to not to store nor leak it, but
// we do not close the process handle to be able to get the exit code
CloseHandle(pi.hThread);
rval->in = _fdopen(in_fd, "w");
rval->out = _fdopen(out_fd, "r");
rval->hProcess = pi.hProcess;
rval->pid = pi.dwProcessId;
return rval;
#else
Expand All @@ -340,6 +345,10 @@ struct child_process *idris2_popen2(char *cmd) {
close(pipes[1]);
return NULL;
}
// make sure no buffers left unflushed before forking
// to save from double-printing
fflush(stdout);
fflush(stderr);
pid_t pid = fork();
if (pid < 0) {
perror("fork");
Expand Down Expand Up @@ -378,6 +387,17 @@ int idris2_popen2ChildPid(struct child_process *ptr) {
return ptr->pid;
}

void *idris2_popen2ChildHandler(struct child_process *ptr) {
#ifdef _WIN32
if (!ptr)
return NULL;
return ptr->hProcess;
#else
// We don't have special child handler in POSIX systems
return NULL;
#endif
}

FILE *idris2_popen2FileIn(struct child_process *ptr) {
if (!ptr)
return NULL;
Expand All @@ -389,3 +409,31 @@ FILE *idris2_popen2FileOut(struct child_process *ptr) {
return NULL;
return ptr->out;
}

int idris2_popen2WaitByHandler(void *hProcess) {
#ifdef _WIN32
DWORD r;
DWORD wres = WaitForSingleObject(hProcess, INFINITE);
IDRIS2_VERIFY(wres == WAIT_OBJECT_0, "waiting after popen2 failed");
int eres = GetExitCodeProcess(hProcess, &r);
IDRIS2_VERIFY(eres != 0, "getting exitcode after popen2 failed");
CloseHandle(hProcess);
return (int)r;
#else
perror("cannot wait by handler in POSIX system");
return -1;
#endif
}

int idris2_popen2WaitByPid(pid_t pid) {
#ifdef _WIN32
perror("cannot wait by pid in Windows");
return -1;
#else
int r = -1;
int w = waitpid(pid, &r, 0);
IDRIS2_VERIFY(w != -1, "waitpid after popen2 failed");
IDRIS2_VERIFY(WIFEXITED(r), "process launched by popen2 didn't exit well");
return WEXITSTATUS(r);
#endif
}
7 changes: 7 additions & 0 deletions support/c/idris_file.h
Original file line number Diff line number Diff line change
Expand Up @@ -3,9 +3,11 @@
#include <stdint.h>
#include <stdio.h>
#include <sys/stat.h>
#include <sys/types.h>

#ifdef _WIN32
#include <io.h>
#include <processthreadsapi.h>
#endif

FILE *idris2_openFile(char *name, char *mode);
Expand Down Expand Up @@ -59,6 +61,11 @@ FILE *idris2_stderr();
struct child_process;

struct child_process *idris2_popen2(char *cmd);

int idris2_popen2ChildPid(struct child_process *ptr);
void *idris2_popen2ChildHandler(struct child_process *ptr);
FILE *idris2_popen2FileIn(struct child_process *ptr);
FILE *idris2_popen2FileOut(struct child_process *ptr);

int idris2_popen2WaitByHandler(void *hProcess);
int idris2_popen2WaitByPid(pid_t pid);
23 changes: 20 additions & 3 deletions tests/allbackends/popen2/Test.idr
Original file line number Diff line number Diff line change
@@ -1,12 +1,29 @@
import System.File

main : IO ()
main = do
Right process <- popen2 "cat"
runPopen2 : (cmd : String) -> IO ()
runPopen2 cmd = do
putStrLn "main thread: before all"
Right process <- popen2 cmd
| Left err => printLn err
printLn $ process.pid > 0
_ <- fPutStrLn process.input "Hello, Idris!\n"
putStrLn "main thread: before closing subinput"
closeFile process.input
putStrLn "main thread: before reading suboutput"
Right result <- fRead process.output
| Left err => printLn err
putStrLn "main thread: before printing suboutput"
putStr result
putStrLn "main thread: before waiting subprocess"
exitCode <- popen2Wait process
putStrLn "subprocess exit code: \{show exitCode}"
putStrLn "main thread: after all"

main : IO ()
main = runPopen2 "cat"

main2 : IO ()
main2 = runPopen2 "cat && exit 4"

main3 : IO ()
main3 = runPopen2 "cat && echo \"Two spaces\" && exit 4"
30 changes: 30 additions & 0 deletions tests/allbackends/popen2/expected
Original file line number Diff line number Diff line change
@@ -1,3 +1,33 @@
main thread: before all
True
main thread: before closing subinput
main thread: before reading suboutput
main thread: before printing suboutput
Hello, Idris!

main thread: before waiting subprocess
subprocess exit code: 0
main thread: after all
------
main thread: before all
True
main thread: before closing subinput
main thread: before reading suboutput
main thread: before printing suboutput
Hello, Idris!

main thread: before waiting subprocess
subprocess exit code: 4
main thread: after all
------
main thread: before all
True
main thread: before closing subinput
main thread: before reading suboutput
main thread: before printing suboutput
Hello, Idris!

Two spaces
main thread: before waiting subprocess
subprocess exit code: 4
main thread: after all
5 changes: 5 additions & 0 deletions tests/allbackends/popen2/run
Original file line number Diff line number Diff line change
@@ -1,3 +1,8 @@
. ../../testutils.sh

run Test.idr
echo "------"
idris2 --exec main2 Test.idr
echo "------"
idris2 --exec main3 Test.idr | sed 's/^"//' | sed 's/" *$//'
# filtering above is to level differences in quotes echoing by Windows and the rest

0 comments on commit 3502f4a

Please sign in to comment.