Remove lamdaRequire

This commit is contained in:
Michael Messer 2020-12-26 18:10:33 -05:00 committed by G. Allais
parent 2a948d1e2e
commit a1f3424ab8
4 changed files with 10 additions and 31 deletions

View File

@ -30,11 +30,10 @@ expression.
.. code-block:: idris
%foreign "node:lambdaRequire:fs:fp=>__require_fs.fstatSync(fp.fd, {bigint: true}).size"
%foreign "node:lambda:fp=>require('fs').fstatSync(fp.fd, {bigint: true}).size"
prim__fileSize : FilePtr -> PrimIO Int
``lambdaRequire`` also accepts a list of separated modules and assigns
them the name ``__require_<module name>``.
``require`` can be used to import javascript modules.
For completion below an example of a foreign available only with ``browser`` codegen:

View File

@ -233,11 +233,11 @@ copyData src start len dest loc
= primIO (prim__copyData src start len dest loc)
%foreign "C:idris2_readBufferData,libidris2_support"
"node:lambdaRequire:fs:(f,b,l,m) => BigInt(__require_fs.readSync(f.fd,b,Number(l), Number(m)))"
"node:lambda:(f,b,l,m) => BigInt(require('fs').readSync(f.fd,b,Number(l), Number(m)))"
prim__readBufferData : FilePtr -> Buffer -> Int -> Int -> PrimIO Int
%foreign "C:idris2_writeBufferData,libidris2_support"
"node:lambdaRequire:fs:(f,b,l,m) => BigInt(__require_fs.writeSync(f.fd,b,Number(l), Number(m)))"
"node:lambda:(f,b,l,m) => BigInt(require('fs').writeSync(f.fd,b,Number(l), Number(m)))"
prim__writeBufferData : FilePtr -> Buffer -> Int -> Int -> PrimIO Int
export

View File

@ -22,7 +22,7 @@ libc fn = "C:" ++ fn ++ ", libc 6"
prim__open : String -> String -> PrimIO FilePtr
%foreign support "idris2_closeFile"
"node:lambdaRequire:fs:(fp) => __require_fs.closeSync(fp.fd)"
"node:lambda:(fp) => require('fs').closeSync(fp.fd)"
prim__close : FilePtr -> PrimIO ()
%foreign support "idris2_fileError"
@ -43,7 +43,7 @@ prim__readChars : Int -> FilePtr -> PrimIO (Ptr String)
prim__readChar : FilePtr -> PrimIO Int
%foreign support "idris2_writeLine"
"node:lambdaRequire:fs:(filePtr, line) => __require_fs.writeSync(filePtr.fd, line, undefined, 'utf-8')"
"node:lambda:(filePtr, line) => require('fs').writeSync(filePtr.fd, line, undefined, 'utf-8')"
prim__writeLine : FilePtr -> String -> PrimIO Int
%foreign support "idris2_eof"
@ -61,7 +61,7 @@ prim__pclose : FilePtr -> PrimIO ()
prim__removeFile : String -> PrimIO Int
%foreign support "idris2_fileSize"
"node:lambdaRequire:fs:fp=>__require_fs.fstatSync(fp.fd, {bigint: true}).size"
"node:lambda:fp=>require('fs').fstatSync(fp.fd, {bigint: true}).size"
prim__fileSize : FilePtr -> PrimIO Int
%foreign support "idris2_fileSize"
@ -71,7 +71,7 @@ prim__fPoll : FilePtr -> PrimIO Int
prim__fileAccessTime : FilePtr -> PrimIO Int
%foreign support "idris2_fileModifiedTime"
"node:lambdaRequire:fs:fp=>__require_fs.fstatSync(fp.fd, {bigint: true}).mtimeMs / 1000n"
"node:lambda:fp=>require('fs').fstatSync(fp.fd, {bigint: true}).mtimeMs / 1000n"
prim__fileModifiedTime : FilePtr -> PrimIO Int
%foreign support "idris2_fileStatusTime"

View File

@ -55,20 +55,6 @@ addConstToPreamble name def =
let v = "const " ++ newName ++ " = (" ++ def ++ ");"
addToPreamble name newName v
requireSafe : String -> String
requireSafe = pack . map (\c => case c of
'@' => '_'
'/' => '_'
'-' => '_'
_ => c
) . unpack
addRequireToPreamble : {auto c : Ref ESs ESSt} -> String -> Core String
addRequireToPreamble name =
do
let newName = "__require_" ++ requireSafe name
let v = "const " ++ newName ++ " = require(" ++ jsString name ++ ");"
addToPreamble name newName v
addSupportToPreamble : {auto c : Ref ESs ESSt} -> String -> String -> Core String
addSupportToPreamble name code =
addToPreamble name name code
@ -333,11 +319,6 @@ makeForeign n x =
let (ty, def) = readCCPart x
case ty of
"lambda" => pure $ "const " ++ jsName n ++ " = (" ++ def ++ ")\n"
"lambdaRequire" =>
do
let (libs, def_) = readCCPart def
traverseList1 addRequireToPreamble (split (==',') libs)
pure $ "const " ++ jsName n ++ " = (" ++ def_ ++ ")\n"
"support" =>
do
let (name, lib_) = break (== ',') def
@ -354,7 +335,7 @@ makeForeign n x =
_ => throw (InternalError $ "invalid string iterator function: " ++ def ++ ", supported functions are \"new\", \"next\"")
_ => throw (InternalError $ "invalid foreign type : " ++ ty ++ ", supported types are \"lambda\", \"lambdaRequire\", \"support\"")
_ => throw (InternalError $ "invalid foreign type : " ++ ty ++ ", supported types are \"lambda\", \"support\"")
foreignDecl : {auto d : Ref Ctxt Defs} -> {auto c : Ref ESs ESSt} -> Name -> List String -> Core String
foreignDecl n ccs =
@ -373,9 +354,8 @@ jsPrim (NS _ (UN "prim__arrayGet")) [_,x,p,_] = pure $ "(" ++ x ++ "[" ++ p ++ "
jsPrim (NS _ (UN "prim__arraySet")) [_,x,p,v,_] = pure $ "(" ++ x ++ "[" ++ p ++ "] = " ++ v ++ ")"
jsPrim (NS _ (UN "prim__os")) [] =
do
os <- addRequireToPreamble "os"
let oscalc = "(o => o === 'linux'?'unix':o==='win32'?'windows':o)"
sysos <- addConstToPreamble "sysos" (oscalc ++ "(" ++ os ++ ".platform())")
sysos <- addConstToPreamble "sysos" (oscalc ++ "(require('os').platform())")
pure sysos
jsPrim (NS _ (UN "void")) [_, _] = jsCrashExp $ jsString $ "Error: Executed 'void'" -- DEPRECATED. TODO: remove when bootstrap has been updated
jsPrim (NS _ (UN "prim__void")) [_, _] = jsCrashExp $ jsString $ "Error: Executed 'void'"