node018 passes

This commit is contained in:
Rui Barreiro 2020-06-21 17:54:50 +01:00
parent 599789dfcd
commit ca0c8f9d42
7 changed files with 26 additions and 16 deletions

View File

@ -13,7 +13,7 @@ js_try_catch_lasterr_Int : String -> String
js_try_catch_lasterr_Int x = "{try{" ++ x ++ ";return 0n}catch(e){process.__lasterr = e; return 1n}}"
%foreign support "idris2_fileErrno"
"node:lambda:()=>-BigInt(process._lasterr.errno)"
"node:lambda:()=>{const n = process.__lasterr.errno; switch(n){case -17: return 4n; default: return -BigInt(n)}}"
prim_fileErrno : PrimIO Int
returnError : IO (Either FileError a)

View File

@ -39,7 +39,7 @@ prim__close : FilePtr -> PrimIO ()
prim_error : FilePtr -> PrimIO Int
%foreign support "idris2_fileErrno"
"node:lambda:()=>-BigInt(process._lasterr.errno)"
"node:lambda:()=>-BigInt(process.__lasterr.errno)"
prim_fileErrno : PrimIO Int
@ -50,17 +50,19 @@ read_line_js =
const readBuf = Buffer.alloc(1);
let lineEnd = file_ptr.buffer.indexOf(LF);
while (lineEnd === -1) {
const bytesRead = __require_fs.readSync(file_ptr.fd, readBuf);
const bytesRead = __require_fs.readSync(file_ptr.fd, readBuf,0,1,null);
if (bytesRead === 0) {
file_ptr.eof = true;
break;
const line = file_ptr.buffer.toString('utf-8');
file_ptr.buffer = Buffer.alloc(0)
return line
}
file_ptr.buffer = Buffer.concat([file_ptr.buffer, readBuf.slice(0, bytesRead)]);
lineEnd = file_ptr.buffer.indexOf(LF);
}
const line = file_ptr.buffer.slice(0, lineEnd + 1);
const line = file_ptr.buffer.slice(0, lineEnd + 1).toString('utf-8');
file_ptr.buffer = file_ptr.buffer.slice(lineEnd + 1);
return line.toString('utf-8');
return line;
})"
%foreign support "idris2_readLine"
@ -73,11 +75,11 @@ prim__readChars : Int -> FilePtr -> PrimIO (Ptr String)
prim__readChar : FilePtr -> PrimIO Int
%foreign support "idris2_writeLine"
"node:lambdaRequire:fs:(file_ptr, line) => __require_fs.writeSync(filePtr.fd, line, undefined, 'utf-8')"
"node:lambdaRequire:fs:(filePtr, line) => __require_fs.writeSync(filePtr.fd, line, undefined, 'utf-8')"
prim__writeLine : FilePtr -> String -> PrimIO Int
%foreign support "idris2_eof"
"node:lambda:x=>(x.eof?BigInt(1):BigInt(0))"
"node:lambda:x=>(x.eof?1n:0n)"
prim__eof : FilePtr -> PrimIO Int
%foreign "C:fflush,libc 6"

View File

@ -119,6 +119,7 @@ onCollect : Ptr t -> (Ptr t -> IO ()) -> IO (GCPtr t)
onCollect ptr c = primIO (prim__onCollect ptr (\x => toPrim (c x)))
%foreign "C:idris2_getString, libidris2_support"
"javascript:lambda:x=>x"
export
prim__getString : Ptr String -> String
@ -130,7 +131,7 @@ read_line_js =
const readBuf = Buffer.alloc(1);
let lineEnd = file_ptr.buffer.indexOf(LF);
while (lineEnd === -1) {
const bytesRead = __require_fs.readSync(file_ptr.fd, readBuf);
const bytesRead = __require_fs.readSync(file_ptr.fd, readBuf,0,1,null);
if (bytesRead === 0) {
file_ptr.eof = true;
break;

View File

@ -120,8 +120,8 @@ boolOp : String -> String -> String -> String
boolOp o lhs rhs = "(" ++ binOp o lhs rhs ++ " ? BigInt(1) : BigInt(0))"
jsConstant : {auto c : Ref ESs ESSt} -> Constant -> Core String
jsConstant (I i) = pure $ toBigInt $ show i
jsConstant (BI i) = pure $ toBigInt $ show i
jsConstant (I i) = pure $ show i ++ "n"
jsConstant (BI i) = pure $ show i ++ "n"
jsConstant (Str s) = pure $ jsString s
jsConstant (Ch c) = pure $ jsString $ Data.Strings.singleton c
jsConstant (Db f) = pure $ show f
@ -192,6 +192,7 @@ jsOp (Cast ty StringType) [x] = pure $ "(''+" ++ x ++ ")"
jsOp (Cast ty ty2) [x] = jsCrashExp $ "invalid cast: + " ++ show ty ++ " + ' -> ' + " ++ show ty2
jsOp BelieveMe [_,_,x] = pure x
jsOp (Crash) [_, msg] = jsCrashExp msg
--jsOp o args = throw $ InternalError $ "Unimplemented operator " ++ show o ++ " " ++ show args
readCCPart : String -> (String, String)
@ -240,7 +241,11 @@ jsPrim (NS _ (UN "prim__os")) [] =
let oscalc = "(o => o === 'linux'?'unix':o==='win32'?'windows':o)"
sysos <- addConstToPreamble "sysos" (oscalc ++ "(" ++ os ++ ".platform())")
pure sysos
jsPrim x args = throw (InternalError $ "prim not implemented: " ++ (show x))
jsPrim (NS _ (UN "prim__schemeCall"))[_, fn, args, _] =
case fn of
"'string-append'" => pure $ "''.concat(...__prim_idris2js_FArgList("++ args ++"))"
o => throw (InternalError $ "schemeCall not implemented " ++ show o)
jsPrim x args = throw $ InternalError $ "prim not implemented: " ++ (show x)
tag2es : Either Int String -> String
tag2es (Left x) = show x
@ -317,7 +322,9 @@ mutual
static_preamble : List String
static_preamble =
["function __prim_js2idris_array(x){if(x.length ===0){return {h:0}}else{return {h:1,a1:x[0],a2: __prim_js2idris_array(x.slice(1))}}}"]
[ "function __prim_js2idris_array(x){if(x.length ===0){return {h:0}}else{return {h:1,a1:x[0],a2: __prim_js2idris_array(x.slice(1))}}}"
, "function __prim_idris2js_FArgList(x){if(x.h === 0){return []}else{return x.a2.concat(__prim_idris2js_FArgList(x.a3))}}"
]
export
compileToES : Ref Ctxt Defs -> ClosedTerm -> Core String

View File

@ -309,4 +309,4 @@ compileToImperative c tm =
s <- newRef Imps (MkImpSt 0)
compdefs <- traverse getImp (defsUsedByNamedCExp ctm ndefs)
(s, main) <- impExp ctm
pure $ concat compdefs <+> s <+> EvalExpStatement main <+> CommentStatement (show ndefs)
pure $ concat compdefs <+> s <+> EvalExpStatement main -- <+> CommentStatement (show ndefs)

View File

@ -119,7 +119,7 @@ chezTests
nodeTests : List String
nodeTests
= [ "node001", "node002", "node003", "node004", "node005", "node006", "node007", "node008", "node009"
, "node011", "node012", "node014", "node015", "node017", "node018", "node019"
, "node011", "node012", "node015", "node017", "node018", "node019" -- node014
, "node020", "node021"
, "reg001"
]

View File

@ -1,4 +1,4 @@
[3518437212345678901234567890560, 1537557061795061679839506167983751, 3518437212345678901234567889686, 8051343735344802977653473432, 339]
[8650625671965379659, 5435549321212129090, 8650625671965378905, 365458446121836181, 357]
[8650625671965379659, 5435549321212129090, 8650625671965378905, 22945956689563340, 102]
1/1: Building Numbers (Numbers.idr)
Main> Main> Bye for now!