Skip to content

Commit

Permalink
[ new ] Add support for bi-directional pipes on POSIX systems (resolves
Browse files Browse the repository at this point in the history
  • Loading branch information
dunhamsteve committed Apr 15, 2023
1 parent a75161c commit 9544162
Show file tree
Hide file tree
Showing 8 changed files with 202 additions and 4 deletions.
2 changes: 2 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -139,6 +139,8 @@
* Changes `getNProcessors` to return the number of online processors rather than
the number of configured processors.

* Adds `popen2` to run a subprocess with bi-directional pipes.

### Contrib

* Adds `Data.List.Sufficient`, a small library defining a structurally inductive view of lists.
Expand Down
61 changes: 57 additions & 4 deletions libs/base/System/File/Process.idr
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ module System.File.Process
import public System.Escape
import public System.File.Error
import public System.File.Mode
import System.FFI
import System.File.Support
import public System.File.Types

Expand All @@ -15,6 +16,20 @@ prim__popen : String -> String -> PrimIO FilePtr
supportNode "pclose"
prim__pclose : FilePtr -> PrimIO Int

data Popen2Result : Type where

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

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

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

%foreign supportC "idris2_popen2FileOut"
prim__popen2FileOut : Ptr Popen2Result -> PrimIO FilePtr

||| Force a write of all user-space buffered data for the given `File`.
|||
||| @ h the file handle to flush
Expand Down Expand Up @@ -42,14 +57,52 @@ popen cmd m = do
then returnError
else pure (Right (FHandle ptr))

namespace Escaped
export
popen : HasIO io => (cmd : List String) -> (m : Mode) -> io (Either FileError File)
popen = popen . escapeCmd

||| Wait for the process associated with the pipe to terminate.
|||
||| @ fh the file handle to the stream to close/wait on
export
pclose : HasIO io => (fh : File) -> io Int
pclose (FHandle h) = primIO (prim__pclose h)

||| Result of a popen2 command, containing the
public export
record SubProcess where
constructor MkSubProcess
||| Process id of the spawned process
pid : Int
||| The input stream of the spawned process
input : File
||| The output stream of the spawned process
output : File

||| Create a new bidirectional pipe by invoking the shell, which is passed the
||| 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.
|||
||| 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.
|||
||| This function is not supported on node at this time.
export
popen2 : HasIO io => (cmd : String) -> io (Either FileError SubProcess)
popen2 cmd = do
ptr <- primIO (prim__popen2 cmd)
if prim__nullPtr ptr /= 0
then returnError
else do
pid <- primIO (prim__popen2ChildPid ptr)
input <- primIO (prim__popen2FileIn ptr)
output <- primIO (prim__popen2FileOut ptr)
free (prim__forgetPtr ptr)
pure $ Right (MkSubProcess pid (FHandle input) (FHandle output))

namespace Escaped
export
popen : HasIO io => (cmd : List String) -> (m : Mode) -> io (Either FileError File)
popen = popen . escapeCmd

export
popen2 : HasIO io => (cmd : List String) -> io (Either FileError SubProcess)
popen2 = popen2 . escapeCmd
117 changes: 117 additions & 0 deletions support/c/idris_file.c
Original file line number Diff line number Diff line change
Expand Up @@ -11,6 +11,7 @@

#ifdef _WIN32
#include "windows/win_utils.h"
#include <windows.h>
#else
#include <sys/select.h>
#include <sys/wait.h>
Expand Down Expand Up @@ -223,3 +224,119 @@ FILE *idris2_stdin() { return stdin; }
FILE *idris2_stdout() { return stdout; }

FILE *idris2_stderr() { return stderr; }

struct child_process {
pid_t pid;
FILE *in;
FILE *out;
};

// Open a bi-directional pipe, returning the above structure
// with pid and two file handles.
struct child_process *idris2_popen2(char *cmd) {
#ifdef _WIN32
SECURITY_ATTRIBUTES saAttr;
HANDLE pipes[4];
STARTUPINFO si;
PROCESS_INFORMATION pi;
ZeroMemory(&pi, sizeof(PROCESS_INFORMATION));

saAttr.nLength = sizeof(SECURITY_ATTRIBUTES);
saAttr.bInheritHandle = TRUE;
saAttr.lpSecurityDescriptor = NULL;

if (!CreatePipe(&pipes[0], &pipes[1], &saAttr, 0) ||
!CreatePipe(&pipes[2], &pipes[3], &saAttr, 0)) {
return NULL;
}
char cmdline[4096];
int len = snprintf(cmdline, 4096, "cmd /c %s", cmd);
if (len > 4095 || len < 0) {
return NULL;
}
ZeroMemory(&si, sizeof(si));
si.cb = sizeof(si);
si.hStdInput = pipes[2];
si.hStdOutput = pipes[1];
// si.hStdError = pipes[1];
si.dwFlags |= STARTF_USESTDHANDLES;
SetHandleInformation(pipes[3], HANDLE_FLAG_INHERIT, 0);
SetHandleInformation(pipes[0], HANDLE_FLAG_INHERIT, 0);
if (!CreateProcess(NULL, cmdline, NULL, NULL, TRUE, 0, NULL, NULL, &si,
&pi)) {
return NULL;
}
struct child_process *rval = malloc(sizeof(struct child_process));
int in_fd = _open_osfhandle((intptr_t)pipes[3], _O_WRONLY);
int out_fd = _open_osfhandle((intptr_t)pipes[0], _O_RDONLY);
CloseHandle(pipes[1]);
CloseHandle(pipes[2]);
CloseHandle(pi.hProcess);
CloseHandle(pi.hThread);
rval->in = _fdopen(in_fd, "w");
rval->out = _fdopen(out_fd, "r");
rval->pid = pi.dwProcessId;
return rval;
#else
int pipes[4];
int err = 0;
err = pipe(&pipes[0]);
if (err) {
return NULL;
}

err = pipe(&pipes[2]);
if (err) {
close(pipes[0]);
close(pipes[1]);
return NULL;
}
pid_t pid = fork();
if (pid < 0) {
perror("fork");
return NULL;
} else if (pid > 0) {
struct child_process *rval = malloc(sizeof(struct child_process));
close(pipes[1]);
close(pipes[2]);
rval->in = fdopen(pipes[3], "w");
rval->out = fdopen(pipes[0], "r");
rval->pid = pid;
return rval;
} else {
close(STDOUT_FILENO);
dup2(pipes[1], STDOUT_FILENO);
close(pipes[0]);
close(pipes[1]);

close(STDIN_FILENO);
dup2(pipes[2], STDIN_FILENO);
close(pipes[2]);
close(pipes[3]);

err = execlp("/bin/sh", "sh", "-c", cmd, NULL);
// We only reach this point if there is an error.
// Maybe report something to stderr so the user knows what's up?
perror("execl");
exit(err);
}
#endif
}

int idris2_popen2ChildPid(struct child_process *ptr) {
if (!ptr)
return 0;
return ptr->pid;
}

FILE *idris2_popen2FileIn(struct child_process *ptr) {
if (!ptr)
return NULL;
return ptr->in;
}

FILE *idris2_popen2FileOut(struct child_process *ptr) {
if (!ptr)
return NULL;
return ptr->out;
}
7 changes: 7 additions & 0 deletions support/c/idris_file.h
Original file line number Diff line number Diff line change
Expand Up @@ -47,3 +47,10 @@ int idris2_fileIsTTY(FILE *f);
FILE *idris2_stdin();
FILE *idris2_stdout();
FILE *idris2_stderr();

struct child_process;

struct child_process *idris2_popen2(char *cmd);
int idris2_popen2ChildPid(struct child_process *ptr);
FILE *idris2_popen2FileIn(struct child_process *ptr);
FILE *idris2_popen2FileOut(struct child_process *ptr);
1 change: 1 addition & 0 deletions tests/Main.idr
Original file line number Diff line number Diff line change
Expand Up @@ -214,6 +214,7 @@ idrisTestsAllBackends cg = MkTestPool
-- RefC implements IEEE standard and distinguishes between 0.0 and -0.0
-- unlike other backends. So turn this test for now.
$ ([ "issue2362" ] <* guard (cg /= C))
++ ([ "popen2" ] <* guard (cg /= Node))
++ [ -- Evaluator
"evaluator004",
-- Unfortunately the behaviour of Double is platform dependent so the
Expand Down
12 changes: 12 additions & 0 deletions tests/allbackends/popen2/Test.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,12 @@
import System.File

main : IO ()
main = do
Right process <- popen2 "cat"
| Left err => printLn err
printLn $ process.pid > 0
_ <- fPutStrLn process.input "Hello, Idris!\n"
closeFile process.input
Right result <- fRead process.output
| Left err => printLn err
putStr result
3 changes: 3 additions & 0 deletions tests/allbackends/popen2/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
True
Hello, Idris!

3 changes: 3 additions & 0 deletions tests/allbackends/popen2/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
rm -rf build

$1 --exec main Test.idr

0 comments on commit 9544162

Please sign in to comment.