mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-19 17:21:59 +03:00
35 lines
788 B
Idris
35 lines
788 B
Idris
||| Magic/Virtual files
|
|
module System.File.Virtual
|
|
|
|
import System.File.Support
|
|
import public System.File.Types
|
|
|
|
%default total
|
|
|
|
%foreign support "idris2_stdin"
|
|
"node:lambda:x=>({fd:0, buffer: Buffer.alloc(0), name:'<stdin>', eof: false})"
|
|
prim__stdin : FilePtr
|
|
|
|
%foreign support "idris2_stdout"
|
|
"node:lambda:x=>({fd:1, buffer: Buffer.alloc(0), name:'<stdout>', eof: false})"
|
|
prim__stdout : FilePtr
|
|
|
|
%foreign support "idris2_stderr"
|
|
"node:lambda:x=>({fd:2, buffer: Buffer.alloc(0), name:'<stderr>', eof: false})"
|
|
prim__stderr : FilePtr
|
|
|
|
||| The standard input.
|
|
export
|
|
stdin : File
|
|
stdin = FHandle prim__stdin
|
|
|
|
||| The standard output.
|
|
export
|
|
stdout : File
|
|
stdout = FHandle prim__stdout
|
|
|
|
||| The standard error.
|
|
export
|
|
stderr : File
|
|
stderr = FHandle prim__stderr
|