System NodeJS additions (#2401)

This commit is contained in:
Mathew Polzin 2022-04-07 02:09:30 -07:00 committed by GitHub
parent 79e733b848
commit c91a768486
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
17 changed files with 121 additions and 49 deletions

View File

@ -10,34 +10,44 @@ public export
DirPtr : Type
DirPtr = AnyPtr
support : String -> String
support fn = "C:" ++ fn ++ ", libidris2_support, idris_directory.h"
||| Shorthand for referring to the C support library
|||
||| @ fn the function name to refer to in the C support library
supportC : (fn : String) -> String
supportC fn = "C:\{fn}, libidris2_support, idris_directory.h"
||| Shorthand for referring to the Node system support library
|||
||| @ fn the function name to refer to in the js/system_support.js file
supportNode : (fn : String) -> String
supportNode fn = "node:support:\{fn},support_system_directory"
ok : HasIO io => a -> io (Either FileError a)
ok x = pure (Right x)
%foreign support "idris2_currentDirectory"
%foreign supportC "idris2_currentDirectory"
"node:lambda:()=>process.cwd()"
prim__currentDir : PrimIO (Ptr String)
%foreign support "idris2_changeDir"
"node:support:changeDir,support_system_directory"
%foreign supportC "idris2_changeDir"
supportNode "changeDir"
prim__changeDir : String -> PrimIO Int
%foreign support "idris2_createDir"
"node:support:createDir,support_system_directory"
%foreign supportC "idris2_createDir"
supportNode "createDir"
prim__createDir : String -> PrimIO Int
%foreign support "idris2_openDir"
%foreign supportC "idris2_openDir"
prim__openDir : String -> PrimIO DirPtr
%foreign support "idris2_closeDir"
%foreign supportC "idris2_closeDir"
prim__closeDir : DirPtr -> PrimIO ()
%foreign support "idris2_removeDir"
%foreign supportC "idris2_removeDir"
supportNode "removeDir"
prim__removeDir : String -> PrimIO ()
%foreign support "idris2_nextDirEntry"
%foreign supportC "idris2_nextDirEntry"
prim__dirEntry : DirPtr -> PrimIO (Ptr String)
||| Data structure for managing the pointer to a directory.
@ -87,6 +97,7 @@ closeDir : HasIO io => Directory -> io ()
closeDir (MkDir d) = primIO (prim__closeDir d)
||| Remove the directory at the specified path.
||| If the directory is not empty, this operation fails.
export
removeDir : HasIO io => String -> io ()
removeDir dirName = primIO (prim__removeDir dirName)

View File

@ -12,11 +12,11 @@ import Data.Buffer
%default total
%foreign support "idris2_readBufferData"
%foreign supportC "idris2_readBufferData"
"node:lambda:(f,b,l,m) => require('fs').readSync(f.fd,b,l,m)"
prim__readBufferData : FilePtr -> Buffer -> (offset : Int) -> (maxbytes : Int) -> PrimIO Int
%foreign support "idris2_writeBufferData"
%foreign supportC "idris2_writeBufferData"
"node:lambda:(f,b,l,m) => require('fs').writeSync(f.fd,b,l,m)"
prim__writeBufferData : FilePtr -> Buffer -> (offset : Int) -> (size : Int) -> PrimIO Int

View File

@ -7,12 +7,12 @@ import public System.File.Types
%default total
%foreign support "idris2_fileError"
%foreign supportC "idris2_fileError"
"node:lambda:x=>(x===1?1:0)"
prim__error : FilePtr -> PrimIO Int
%foreign support "idris2_fileErrno"
"node:support:fileErrno,support_system_file"
%foreign supportC "idris2_fileErrno"
supportNode "fileErrno"
prim__fileErrno : PrimIO Int
||| The types of errors that can occur during file operations.

View File

@ -7,11 +7,11 @@ import public System.File.Types
%default total
%foreign support "idris2_openFile"
"node:support:openFile,support_system_file"
%foreign supportC "idris2_openFile"
supportNode "openFile"
prim__open : String -> String -> PrimIO FilePtr
%foreign support "idris2_closeFile"
%foreign supportC "idris2_closeFile"
"node:lambda:(fp) => require('fs').closeSync(fp.fd)"
prim__close : FilePtr -> PrimIO ()

View File

@ -7,24 +7,24 @@ import public System.File.Types
%default total
%foreign support "idris2_fileSize"
%foreign supportC "idris2_fileSize"
"node:lambda:fp=>require('fs').fstatSync(fp.fd).size"
prim__fileSize : FilePtr -> PrimIO Int
%foreign support "idris2_fileSize"
%foreign supportC "idris2_fileSize"
prim__fPoll : FilePtr -> PrimIO Int
%foreign support "idris2_fileAccessTime"
%foreign supportC "idris2_fileAccessTime"
prim__fileAccessTime : FilePtr -> PrimIO Int
%foreign support "idris2_fileModifiedTime"
%foreign supportC "idris2_fileModifiedTime"
"node:lambda:fp=>require('fs').fstatSync(fp.fd).mtimeMs / 1000"
prim__fileModifiedTime : FilePtr -> PrimIO Int
%foreign support "idris2_fileStatusTime"
%foreign supportC "idris2_fileStatusTime"
prim__fileStatusTime : FilePtr -> PrimIO Int
%foreign support "idris2_fileIsTTY"
%foreign supportC "idris2_fileIsTTY"
"node:lambda:fp=>Number(require('tty').isatty(fp.fd))"
prim__fileIsTTY : FilePtr -> PrimIO Int

View File

@ -6,8 +6,8 @@ import public System.File.Types
%default total
%foreign support "idris2_chmod"
"node:support:chmod,support_system_file"
%foreign supportC "idris2_chmod"
supportNode "chmod"
prim__chmod : String -> Int -> PrimIO Int
||| The UNIX file modes.

View File

@ -8,9 +8,9 @@ import public System.File.Types
%foreign "C:fflush,libc 6"
prim__flush : FilePtr -> PrimIO Int
%foreign support "idris2_popen"
%foreign supportC "idris2_popen"
prim__popen : String -> String -> PrimIO FilePtr
%foreign support "idris2_pclose"
%foreign supportC "idris2_pclose"
prim__pclose : FilePtr -> PrimIO Int
||| Force a write of all user-space buffered data for the given `File`.

View File

@ -11,28 +11,29 @@ import public System.File.Types
%default total
%foreign support "idris2_seekLine"
"node:support:seekLine,support_system_file"
%foreign supportC "idris2_seekLine"
supportNode "seekLine"
prim__seekLine : FilePtr -> PrimIO Int
%foreign support "idris2_readLine"
"node:support:readLine,support_system_file"
%foreign supportC "idris2_readLine"
supportNode "readLine"
prim__readLine : FilePtr -> PrimIO (Ptr String)
%foreign support "idris2_readChars"
%foreign supportC "idris2_readChars"
prim__readChars : Int -> FilePtr -> PrimIO (Ptr String)
%foreign "C:fgetc,libc 6"
prim__readChar : FilePtr -> PrimIO Int
%foreign support "idris2_writeLine"
%foreign supportC "idris2_writeLine"
"node:lambda:(filePtr, line) => require('fs').writeSync(filePtr.fd, line, undefined, 'utf-8')"
prim__writeLine : FilePtr -> String -> PrimIO Int
%foreign support "idris2_eof"
%foreign supportC "idris2_eof"
"node:lambda:x=>(x.eof?1:0)"
prim__eof : FilePtr -> PrimIO Int
%foreign support "idris2_removeFile"
%foreign supportC "idris2_removeFile"
supportNode "removeFile"
prim__removeFile : String -> PrimIO Int
||| Seek through the next newline.

View File

@ -2,13 +2,19 @@ module System.File.Support
%default total
||| Shorthand for a function in the C support libary
||| (libidris2_support, idris_file.h)
||| Shorthand for referring to the C support library
|||
||| @ fn the function name to refer to in the C support library
public export
support : String -> String
support fn = "C:" ++ fn ++ ", libidris2_support, idris_file.h"
supportC : (fn : String) -> String
supportC fn = "C:\{fn}, libidris2_support, idris_file.h"
||| Shorthand for referring to the Node system support library
|||
||| @ fn the function name to refer to in the js/system_support_file.js file
public export
supportNode : (fn : String) -> String
supportNode fn = "node:support:\{fn},support_system_file"
||| Wrap x in the `Right` part of an `io . Either`.
export

View File

@ -6,15 +6,15 @@ import public System.File.Types
%default total
%foreign support "idris2_stdin"
%foreign supportC "idris2_stdin"
"node:lambda:x=>({fd:0, buffer: Buffer.alloc(0), name:'<stdin>', eof: false})"
prim__stdin : FilePtr
%foreign support "idris2_stdout"
%foreign supportC "idris2_stdout"
"node:lambda:x=>({fd:1, buffer: Buffer.alloc(0), name:'<stdout>', eof: false})"
prim__stdout : FilePtr
%foreign support "idris2_stderr"
%foreign supportC "idris2_stderr"
"node:lambda:x=>({fd:2, buffer: Buffer.alloc(0), name:'<stderr>', eof: false})"
prim__stderr : FilePtr

View File

@ -23,6 +23,7 @@ isWindows : Bool
isWindows = os `elem` ["windows", "mingw32", "cygwin32"]
%foreign "C:idris2_getNProcessors, libidris2_support, idris_support.h"
"node:lambda:() => require('os').cpus().length"
prim__getNProcessors : PrimIO Int
||| Get the number of processors on the system. Returns `Nothing` if we somehow

View File

@ -12,11 +12,17 @@ import Data.Buffer
%default total
%foreign support "idris2_readBufferData"
support : String -> String
support fn = "C:\{fn}, libidris2_support, idris_file.h"
-- ^ Only needed until version following 0.5.1 is released at which point
-- any code in compiler src can be refactored to use renamed @supportC@
-- function.
%foreign Buffer.support "idris2_readBufferData"
"node:lambda:(f,b,l,m) => require('fs').readSync(f.fd,b,l,m)"
prim__readBufferData : FilePtr -> Buffer -> (offset : Int) -> (maxbytes : Int) -> PrimIO Int
%foreign support "idris2_writeBufferData"
%foreign Buffer.support "idris2_writeBufferData"
"node:lambda:(f,b,l,m) => require('fs').writeSync(f.fd,b,l,m)"
prim__writeBufferData : FilePtr -> Buffer -> (offset : Int) -> (size : Int) -> PrimIO Int

View File

@ -2,10 +2,10 @@ const support_system_directory_fs = require("fs");
function support_system_directory_changeDir(d){
try{
process.chdir(d);
process.chdir(d)
return 0
}catch(e){
process.__lasterr = e;
process.__lasterr = e
return 1
}
}
@ -15,7 +15,17 @@ function support_system_directory_createDir(d){
support_system_directory_fs.mkdirSync(d)
return 0
}catch(e){
process.__lasterr = e;
process.__lasterr = e
return 1
}
}
function support_system_directory_removeDir(d){
try{
support_system_directory_fs.rmdirSync(d)
return 0
}catch(e){
process.__lasterr = e
return 1
}
}

View File

@ -81,3 +81,13 @@ function support_system_file_chmod (filename, mode) {
return 1
}
}
function support_system_file_removeFile (filename) {
try {
support_system_file_fs.unlinkSync(filename)
return 0
} catch (e) {
process.__lasterr = e
return 1
}
}

View File

@ -0,0 +1,14 @@
import System
import System.File
main : IO ()
main = do
Right () <- writeFile "hello.txt" "hello"
| _ => die "Failed to create file during test setup"
True <- exists "hello.txt"
| _ => die "assumption that written file exists did not hold."
Right () <- removeFile "hello.txt"
| _ => die "Failed to remove file"
case !(exists "hello.txt") of
True => die "Supposedly removed file still exists."
False => putStrLn "Good deal!"

View File

@ -0,0 +1,2 @@
Good deal!
Good deal!

View File

@ -0,0 +1,11 @@
rm -rf build
# @removeFile@ uses primitive functions with definitions for both
# C (supported by most backends) and Node.
$1 --cg chez -o test Test.idr
./build/exec/test
rm -rf ./build
$1 --cg node -o test.js Test.idr
node ./build/exec/test.js