[base] add missing node ffi functions (#2427)

This commit is contained in:
Zoe Stafford 2022-04-22 15:45:52 +01:00 committed by GitHub
parent 0010768ee7
commit 68bcacf3ec
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
3 changed files with 38 additions and 0 deletions

View File

@ -117,6 +117,7 @@ getBits32 buf offset
= primIO (prim__getBits32 buf offset)
%foreign "scheme:blodwen-buffer-setbits64"
"node:lambda:(buf,offset,value)=>buf.writeBigUInt64LE(value, offset)"
prim__setBits64 : Buffer -> Int -> Bits64 -> PrimIO ()
export %inline
@ -125,6 +126,7 @@ setBits64 buf offset val
= primIO (prim__setBits64 buf offset val)
%foreign "scheme:blodwen-buffer-getbits64"
"node:lambda:(buf,offset)=>buf.readBigUInt64LE(offset)"
prim__getBits64 : Buffer -> (offset : Int) -> PrimIO Bits64
export %inline
@ -194,6 +196,7 @@ getDouble buf offset
export
%foreign "scheme:blodwen-stringbytelen"
"C:strlen, libc 6"
"javascript:lambda:(string)=>new TextEncoder().encode(string).length"
stringByteLength : String -> Int
%foreign "scheme:blodwen-buffer-setstring"

View File

@ -38,9 +38,11 @@ prim__changeDir : String -> PrimIO Int
prim__createDir : String -> PrimIO Int
%foreign supportC "idris2_openDir"
supportNode "openDir"
prim__openDir : String -> PrimIO DirPtr
%foreign supportC "idris2_closeDir"
supportNode "closeDir"
prim__closeDir : DirPtr -> PrimIO ()
%foreign supportC "idris2_removeDir"
@ -48,6 +50,7 @@ prim__closeDir : DirPtr -> PrimIO ()
prim__removeDir : String -> PrimIO ()
%foreign supportC "idris2_nextDirEntry"
supportNode "dirEntry"
prim__dirEntry : DirPtr -> PrimIO (Ptr String)
||| Data structure for managing the pointer to a directory.

View File

@ -29,3 +29,35 @@ function support_system_directory_removeDir(d){
return 1
}
}
function support_system_directory_openDir(d) {
try{
return support_system_directory_fs.opendirSync(d)
}catch(e){
process.__lasterr = e
return null
}
}
function support_system_directory_closeDir(d) {
try{
d.closeSync()
}catch(e){
process.__lasterr = e
return null
}
}
function support_system_directory_dirEntry(d) {
try{
const dir = d.readSync()
if (dir == null) {
return null
} else {
return dir.name
}
}catch(e){
process.__lasterr = e
return null
}
}