moved big foreign functions to support and added outputDir

This commit is contained in:
Rui Barreiro 2020-07-06 16:58:02 +01:00
parent b3216758e9
commit 68c9990c8a
10 changed files with 108 additions and 71 deletions

View File

@ -119,9 +119,11 @@ install-support:
mkdir -p ${PREFIX}/idris2-${IDRIS2_VERSION}/support/chez
mkdir -p ${PREFIX}/idris2-${IDRIS2_VERSION}/support/racket
mkdir -p ${PREFIX}/idris2-${IDRIS2_VERSION}/support/gambit
mkdir -p ${PREFIX}/idris2-${IDRIS2_VERSION}/support/js
install support/chez/* ${PREFIX}/idris2-${IDRIS2_VERSION}/support/chez
install support/racket/* ${PREFIX}/idris2-${IDRIS2_VERSION}/support/racket
install support/gambit/* ${PREFIX}/idris2-${IDRIS2_VERSION}/support/gambit
install support/js/* ${PREFIX}/idris2-${IDRIS2_VERSION}/support/js
@${MAKE} -C support/c install
install-libs:

View File

@ -9,11 +9,8 @@ DirPtr = AnyPtr
support : String -> String
support fn = "C:" ++ fn ++ ", libidris2_support"
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:()=>{const n = process.__lasterr.errno; switch(n){case -17: return 4n; default: return -BigInt(n)}}"
"node:support:fileErrno,support_system_directory"
prim_fileErrno : PrimIO Int
returnError : HasIO io => io (Either FileError a)
@ -35,11 +32,11 @@ ok x = pure (Right x)
prim_currentDir : PrimIO (Ptr String)
%foreign support "idris2_changeDir"
("node:lambda:d=>" ++ js_try_catch_lasterr_Int "process.chdir(d)")
"node:support:changeDir,support_system_directory"
prim_changeDir : String -> PrimIO Int
%foreign support "idris2_createDir"
("node:lambdaRequire:fs:d=>" ++ js_try_catch_lasterr_Int "__require_fs.mkdirSync(d)")
"node:support:createDir,support_system_directory"
prim_createDir : String -> PrimIO Int
%foreign support "idris2_openDir"

View File

@ -17,21 +17,12 @@ support fn = "C:" ++ fn ++ ", libidris2_support"
libc : String -> String
libc fn = "C:" ++ fn ++ ", libc 6"
js_try_catch_lasterr_Int : String -> String
js_try_catch_lasterr_Int x = "{try{" ++ x ++ ";return 0n}catch(e){process.__lasterr = e; return 1n}}"
js_try_catch_lasterr_Ptr : String -> String
js_try_catch_lasterr_Ptr x = "{try{" ++ x ++ "}catch(e){process.__lasterr = e; return null}}"
js_open_file : String
js_open_file = "return {fd:__require_fs.openSync(n, m), buffer: Buffer.alloc(0), name:n, eof: false}"
%foreign support "idris2_openFile"
("node:lambdaRequire:fs:(n, m) =>" ++ js_try_catch_lasterr_Ptr js_open_file)
"node:support:openFile,support_system_file"
prim__open : String -> String -> Int -> PrimIO FilePtr
%foreign support "idris2_closeFile"
("node:lambdaRequire:fs:(fp) => __require_fs.closeSync(fp.fd)")
"node:lambdaRequire:fs:(fp) => __require_fs.closeSync(fp.fd)"
prim__close : FilePtr -> PrimIO ()
%foreign support "idris2_fileError"
@ -42,31 +33,8 @@ prim_error : FilePtr -> PrimIO Int
"node:lambda:()=>-BigInt(process.__lasterr.errno)"
prim_fileErrno : PrimIO Int
read_line_js : String
read_line_js =
"(file_ptr =>{
const LF = 0x0a;
const readBuf = Buffer.alloc(1);
let lineEnd = file_ptr.buffer.indexOf(LF);
while (lineEnd === -1) {
const bytesRead = __require_fs.readSync(file_ptr.fd, readBuf,0,1,null);
if (bytesRead === 0) {
file_ptr.eof = true;
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).toString('utf-8');
file_ptr.buffer = file_ptr.buffer.slice(lineEnd + 1);
return line;
})"
%foreign support "idris2_readLine"
("node:lambda:" ++ read_line_js)
"node:support:readLine,support_system_file"
prim__readLine : FilePtr -> PrimIO (Ptr String)
%foreign support "idris2_readChars"
@ -122,7 +90,7 @@ prim__stdout : FilePtr
prim__stderr : FilePtr
%foreign libc "chmod"
("node:lambdaRequire:fs:(filename, mode) => " ++ js_try_catch_lasterr_Int "__require_fs.chmodSync(filename, Number(mode))")
"node:support:chmod,support_system_file"
prim__chmod : String -> Int -> PrimIO Int
modeStr : Mode -> String

View File

@ -1510,28 +1510,8 @@ prim__putChar : Char -> (1 x : %World) -> IORes ()
%foreign "C:getchar,libc 6"
%extern prim__getChar : (1 x : %World) -> IORes Char
read_line_js : String
read_line_js =
"node:lambdaRequire:fs:(file_ptr =>{
const LF = 0x0a;
const readBuf = Buffer.alloc(1);
let lineEnd = file_ptr.buffer.indexOf(LF);
while (lineEnd === -1) {
const bytesRead = __require_fs.readSync(file_ptr.fd, readBuf,0,1,null);
if (bytesRead === 0) {
file_ptr.eof = true;
break;
}
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);
file_ptr.buffer = file_ptr.buffer.slice(lineEnd + 1);
return line.toString('utf-8');
})({fd:0, buffer: Buffer.alloc(0), name:'<stdin>', eof: false})"
%foreign "C:idris2_getStr,libidris2_support"
read_line_js
"node:support:getStr,support_system_file"
prim__getStr : PrimIO String
%foreign "C:idris2_putStr,libidris2_support"

View File

@ -6,6 +6,8 @@ import Data.Strings
import Data.SortedMap
import Data.String.Extra
import Core.Directory
data ESs : Type where
record ESSt where
@ -59,6 +61,11 @@ addRequireToPreamble 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
jsIdent : String -> String
jsIdent s = concatMap okchar (unpack s)
where
@ -256,7 +263,7 @@ searchForeign prefixes (x::xs) =
else searchForeign prefixes xs
makeForeign : {auto c : Ref ESs ESSt} -> Name -> String -> Core String
makeForeign : {auto d : Ref Ctxt Defs} -> {auto c : Ref ESs ESSt} -> Name -> String -> Core String
makeForeign n x =
do
let (ty, def) = readCCPart x
@ -267,9 +274,18 @@ makeForeign n x =
let (libs, def_) = readCCPart def
traverse addRequireToPreamble (split (==',') libs)
pure $ "const " ++ jsName n ++ " = (" ++ def_ ++ ")\n"
"support" =>
do
let (name, lib_) = break (== ',') def
let lib = drop 1 lib_
lib_code <- readDataFile ("js/" ++ lib ++ ".js")
addSupportToPreamble lib lib_code
pure $ "const " ++ jsName n ++ " = " ++ lib ++ "_" ++ name ++ "\n"
_ => throw (InternalError $ "invalid foreign type : " ++ ty ++ ", supporte types are lambda")
foreignDecl : {auto c : Ref ESs ESSt} -> Name -> List String -> Core String
foreignDecl : {auto d : Ref Ctxt Defs} -> {auto c : Ref ESs ESSt} -> Name -> List String -> Core String
foreignDecl n ccs =
do
s <- get ESs
@ -301,7 +317,7 @@ tag2es (Left x) = show x
tag2es (Right x) = jsString x
mutual
impExp2es : {auto c : Ref ESs ESSt} -> ImperativeExp -> Core String
impExp2es : {auto d : Ref Ctxt Defs} -> {auto c : Ref ESs ESSt} -> ImperativeExp -> Core String
impExp2es (IEVar n) =
pure $ jsName n
impExp2es (IELambda args body) =
@ -330,7 +346,7 @@ mutual
impExp2es IENull =
pure "undefined"
imperative2es : {auto c : Ref ESs ESSt} -> Nat -> ImperativeStatement -> Core String
imperative2es : {auto d : Ref Ctxt Defs} -> {auto c : Ref ESs ESSt} -> Nat -> ImperativeStatement -> Core String
imperative2es indent DoNothing =
pure ""
imperative2es indent (SeqStatement x y) =
@ -367,7 +383,7 @@ mutual
imperative2es indent (ForEverLoop x) =
pure $ nSpaces indent ++ "while(true){\n" ++ !(imperative2es (indent+1) x) ++ "\n" ++ nSpaces indent ++ "}"
alt2es : {auto c : Ref ESs ESSt} -> Nat -> (ImperativeExp, ImperativeStatement) -> Core String
alt2es : {auto d : Ref Ctxt Defs} -> {auto c : Ref ESs ESSt} -> Nat -> (ImperativeExp, ImperativeStatement) -> Core String
alt2es indent (e, b) = pure $ nSpaces indent ++ "case " ++ !(impExp2es e) ++ ":\n" ++
!(imperative2es (indent+1) b) ++ "\n" ++ nSpaces (indent+1) ++ "break;\n"

View File

@ -46,7 +46,7 @@ compileExpr : Ref Ctxt Defs -> (tmpDir : String) -> (outputDir : String) ->
compileExpr c tmpDir outputDir tm outfile
= do es <- compileToJS c tm
let res = if toLower (takeLast 5 outfile) == ".html" then htmlHeader ++ es ++ htmlFooter else es
let out = outfile
let out = outputDir </> outfile
Right () <- coreLift (writeFile out res)
| Left err => throw (FileErr out err)
pure (Just out)

View File

@ -30,7 +30,7 @@ compileExpr : Ref Ctxt Defs -> (tmpDir : String) -> (outputDir : String) ->
ClosedTerm -> (outfile : String) -> Core (Maybe String)
compileExpr c tmpDir outputDir tm outfile
= do es <- compileToNode c tm
let out = outfile
let out = outputDir </> outfile
Right () <- coreLift (writeFile out es)
| Left err => throw (FileErr out err)
pure (Just out)

View File

@ -1,7 +1,6 @@
module Compiler.ES.TailRec
import Compiler.ES.ImperativeAst
import Debug.Trace
hasTailCall : Name -> ImperativeStatement -> Bool
hasTailCall n (SeqStatement x y) =

View File

@ -0,0 +1,29 @@
const support_system_directory_fs = require("fs")
function support_system_directory_fileErrno(){
const n = process.__lasterr.errno;
switch(n){
case -17: return 4n;
default: return -BigInt(n)
}
}
function support_system_directory_changeDir(d){
try{
process.chdir(d);
return 0n
}catch(e){
process.__lasterr = e;
return 1n
}
}
function support_system_directory_createDir(d){
try{
support_system_directory_fs.mkdirSync(d)
return 0n
}catch(e){
process.__lasterr = e;
return 1n
}
}

View File

@ -0,0 +1,46 @@
const support_system_file_fs = require("fs")
function support_system_file_readLine(file_ptr){
const LF = 0x0a;
const readBuf = Buffer.alloc(1);
let lineEnd = file_ptr.buffer.indexOf(LF);
while (lineEnd === -1) {
const bytesRead = support_system_file_fs.readSync(file_ptr.fd, readBuf,0,1,null);
if (bytesRead === 0) {
file_ptr.eof = true;
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).toString('utf-8');
file_ptr.buffer = file_ptr.buffer.slice(lineEnd + 1);
return line;
}
function support_system_file_getStr(){
return support_system_file_readLine({fd:0, buffer: Buffer.alloc(0), name:'<stdin>', eof: false})
}
function support_system_file_openFile(n,m){
try{
const fd = support_system_file_fs.openSync(n, m)
return {fd: fd, buffer: Buffer.alloc(0), name:n, eof: false}
}catch(e){
process.__lasterr = e; return null
}
}
function support_system_file_chmod(filename, mode){
try{
support_system_file_fs.chmodSync(filename, Number(mode))
return 0n
}catch(e){
process.__lasterr = e;
return 1n
}
}