mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 20:51:43 +03:00
Merge pull request #2914 from dunhamsteve/enableRawMode
[ base ] Support unbuffered reading of stdin
This commit is contained in:
commit
79116320dd
@ -122,6 +122,8 @@
|
||||
* Adds `Vect.allFins` for generating all the `Fin` elements as a `Vect` with
|
||||
matching length to the number of elements.
|
||||
|
||||
* Add `withRawMode`, `enableRawMode`, `resetRawMode` for character at a time input on stdin.
|
||||
|
||||
#### System
|
||||
|
||||
* Changes `getNProcessors` to return the number of online processors rather than
|
||||
|
@ -98,6 +98,47 @@ prim__setEnv : String -> String -> Int -> PrimIO Int
|
||||
supportNode "unsetEnv"
|
||||
prim__unsetEnv : String -> PrimIO Int
|
||||
|
||||
%foreign "C:idris2_enableRawMode, libidris2_support, idris_support.h"
|
||||
prim__enableRawMode : (1 x : %World) -> IORes Int
|
||||
|
||||
%foreign "C:idris2_resetRawMode, libidris2_support, idris_support.h"
|
||||
prim__resetRawMode : (1 x : %World) -> IORes ()
|
||||
|
||||
||| `enableRawMode` enables raw mode for stdin, allowing characters
|
||||
||| to be read one at a time, without buffering or echoing.
|
||||
||| If `enableRawMode` is used, the program should call `resetRawMode` before
|
||||
||| exiting. Consider using `withRawMode` instead to ensure the tty is reset.
|
||||
|||
|
||||
||| This is not supported on windows.
|
||||
export
|
||||
enableRawMode : HasIO io => io (Either FileError ())
|
||||
enableRawMode =
|
||||
case !(primIO prim__enableRawMode) of
|
||||
0 => pure $ Right ()
|
||||
_ => returnError
|
||||
|
||||
||| `resetRawMode` resets stdin raw mode to original state if
|
||||
||| `enableRawMode` had been previously called.
|
||||
export
|
||||
resetRawMode : HasIO io => io ()
|
||||
resetRawMode = primIO prim__resetRawMode
|
||||
|
||||
||| `withRawMode` performs a given operation after setting stdin to raw mode
|
||||
||| and ensure that stdin is reset to its original state afterwards.
|
||||
|||
|
||||
||| This is not supported on windows.
|
||||
export
|
||||
withRawMode : HasIO io =>
|
||||
(onError : FileError -> io a) ->
|
||||
(onSuccess : () -> io a) ->
|
||||
io a
|
||||
withRawMode onError onSuccess = do
|
||||
Right () <- enableRawMode
|
||||
| Left err => onError err
|
||||
result <- onSuccess ()
|
||||
resetRawMode
|
||||
pure result
|
||||
|
||||
||| Retrieve the specified environment variable's value string, or `Nothing` if
|
||||
||| there is no such environment variable.
|
||||
|||
|
||||
|
@ -5,6 +5,9 @@
|
||||
#include <stdio.h>
|
||||
#include <stdlib.h>
|
||||
#include <string.h>
|
||||
#ifndef _WIN32
|
||||
#include <termios.h>
|
||||
#endif
|
||||
#include <time.h>
|
||||
#include <unistd.h>
|
||||
|
||||
@ -98,6 +101,36 @@ int idris2_setenv(const char *name, const char *value, int overwrite) {
|
||||
#endif
|
||||
}
|
||||
|
||||
#ifndef _WIN32
|
||||
// The initial state of stdin is stored here if enableRawMode is called.
|
||||
struct termios *initial_termios = NULL;
|
||||
#endif
|
||||
|
||||
int idris2_enableRawMode() {
|
||||
#ifndef _WIN32
|
||||
struct termios ti;
|
||||
int rval = tcgetattr(STDIN_FILENO, &ti);
|
||||
if (rval != 0)
|
||||
return rval;
|
||||
if (!initial_termios) {
|
||||
initial_termios = malloc(sizeof(struct termios));
|
||||
*initial_termios = ti;
|
||||
}
|
||||
ti.c_lflag &= ~(ECHO | ICANON);
|
||||
return tcsetattr(STDIN_FILENO, TCSAFLUSH, &ti);
|
||||
#else
|
||||
return -1;
|
||||
#endif
|
||||
}
|
||||
|
||||
void idris2_resetRawMode() {
|
||||
#ifndef _WIN32
|
||||
if (initial_termios != NULL) {
|
||||
tcsetattr(STDIN_FILENO, TCSAFLUSH, initial_termios);
|
||||
}
|
||||
#endif
|
||||
}
|
||||
|
||||
int idris2_unsetenv(const char *name) {
|
||||
#ifdef _WIN32
|
||||
return win32_modenv(name, "", 1);
|
||||
|
Loading…
Reference in New Issue
Block a user