Skip to content

Commit

Permalink
Merge pull request #4073 from rbarreiro/exportiface_js
Browse files Browse the repository at this point in the history
--interface for node backend
  • Loading branch information
melted committed Sep 20, 2017
2 parents bd8ea0a + bea08ac commit 6d954d8
Show file tree
Hide file tree
Showing 12 changed files with 113 additions and 18 deletions.
1 change: 1 addition & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -15,6 +15,7 @@
## Tool Updates
+ Private functions are no longer visible in the REPL except for modules
that are explicitly loaded.
+ The --interface option now creates CommonJS modules on the node backend.

# New in 1.1.1

Expand Down
1 change: 1 addition & 0 deletions CONTRIBUTORS
Original file line number Diff line number Diff line change
Expand Up @@ -177,6 +177,7 @@ Ricky Elrod
Robert Clipsham
Robert Edström
Ross Meikleham
Rui Barreiro
Ryan Scott
Sam Elliott
Sam T
Expand Down
14 changes: 11 additions & 3 deletions codegen/idris-codegen-node/Main.hs
Original file line number Diff line number Diff line change
Expand Up @@ -7,10 +7,15 @@ import Idris.Options
import IRTS.CodegenJavaScript
import IRTS.Compiler

import Paths_idris

import Control.Monad

import System.Environment
import System.Exit

data Opts = Opts { inputs :: [FilePath]
, interface :: Bool
, output :: FilePath
}

Expand All @@ -20,17 +25,20 @@ showUsage = do putStrLn "A code generator which is intended to be called by the

getOpts :: IO Opts
getOpts = do xs <- getArgs
return $ process (Opts [] "main.js") xs
return $ process (Opts [] False "main.js") xs
where
process opts ("-o":o:xs) = process (opts { output = o }) xs
process opts ("--interface":xs) = process (opts { interface = True }) xs
process opts (x:xs) = process (opts { inputs = x:inputs opts }) xs
process opts [] = opts

jsMain :: Opts -> Idris ()
jsMain opts = do elabPrims
loadInputs (inputs opts) Nothing
mainProg <- elabMain
ir <- compile (Via IBCFormat "node") (output opts) (Just mainProg)
mainProg <- if interface opts
then return Nothing
else liftM Just elabMain
ir <- compile (Via IBCFormat "node") (output opts) mainProg
runIO $ codegenNode ir

main :: IO ()
Expand Down
21 changes: 15 additions & 6 deletions src/IRTS/CodegenJavaScript.hs
Original file line number Diff line number Diff line change
Expand Up @@ -55,15 +55,24 @@ codegenJavaScript ci =

codegenNode :: CodeGenerator
codegenNode ci =
do
codegenJs (CGConf { header = "#!/usr/bin/env node\n"
if interfaces ci then
codegenJs (CGConf { header = ""
, footer = ""
, jsbnPath = "jsbn/jsbn-browser.js"
, extraRunTime = "Runtime-node.js"
}
)
ci
setPermissions (outputFile ci) (emptyPermissions { readable = True
, executable = True
, writable = True
})
else
do
codegenJs (CGConf { header = "#!/usr/bin/env node\n"
, footer = ""
, jsbnPath = "jsbn/jsbn-browser.js"
, extraRunTime = "Runtime-node.js"
}
)
ci
setPermissions (outputFile ci) (emptyPermissions { readable = True
, executable = True
, writable = True
})
2 changes: 1 addition & 1 deletion src/IRTS/Compiler.hs
Original file line number Diff line number Diff line change
Expand Up @@ -142,7 +142,7 @@ generate codegen mainmod ir
return fn
let cmd = "idris-codegen-" ++ cg
args = [input, "-o", outputFile ir] ++ compilerFlags ir
exit <- rawSystem cmd args
exit <- rawSystem cmd (if interfaces ir then "--interface" : args else args)
when (exit /= ExitSuccess) $
putStrLn ("FAILURE: " ++ show cmd ++ " " ++ show args)
Bytecode -> dumpBC (simpleDecls ir) (outputFile ir)
Expand Down
45 changes: 39 additions & 6 deletions src/IRTS/JavaScript/Codegen.hs
Original file line number Diff line number Diff line change
Expand Up @@ -14,6 +14,7 @@ module IRTS.JavaScript.Codegen( codegenJs

import Idris.Core.TT
import IRTS.CodegenCommon
import IRTS.Exports
import IRTS.JavaScript.AST
import IRTS.JavaScript.LangTransforms
import IRTS.JavaScript.Name
Expand All @@ -39,7 +40,6 @@ import System.Directory (doesFileExist)
import System.Environment
import System.FilePath


-- | Code generation stats hold information about the generated user
-- code. Based on that information we add additional code to make
-- things work.
Expand Down Expand Up @@ -98,14 +98,42 @@ isYes (Just "Y") = True
isYes (Just "y") = True
isYes _ = False

makeExportDecls :: Map Name LDecl -> ExportIFace -> [Text]
makeExportDecls defs (Export _ _ e) =
concatMap makeExport e
where
uncurryF name argTy (Just args) =
if length argTy == length args then name
else T.concat [ "function(){ return "
, name
, ".apply(this, Array.prototype.slice.call(arguments, 0,", T.pack $ show $ length args,"))"
, T.concat $ map (\x -> T.concat ["(arguments[", T.pack $ show x , "])"]) [length args .. (length argTy - 1)]
, "}"
]
uncurryF name argTy Nothing = name

addApplyForIO (FIO _) f = T.concat ["function(){ return (",f, ").apply(this,arguments)()}"]
addApplyForIO _ f = f

makeExport (ExportData _) =
[]
makeExport (ExportFun name (FStr exportname) retTy argTy) =
[T.concat [ T.pack $ exportname
, ": "
, addApplyForIO retTy $ uncurryF (jsName name) argTy (getArgList' name defs)
]
]

codegenJs :: CGConf -> CodeGenerator
codegenJs conf ci =
do
debug <- isYes <$> lookupEnv "IDRISJS_DEBUG"
let defs' = Map.fromList $ liftDecls ci
let defs = globlToCon defs'
let used = Map.elems $ removeDeadCode defs [sMN 0 "runMain"]
let iface = interfaces ci
let used = if iface then
Map.elems $ removeDeadCode defs (getExpNames $ exportDecls ci)
else Map.elems $ removeDeadCode defs [sMN 0 "runMain"]
when debug $ do
writeFile (outputFile ci ++ ".LDeclsDebug") $ (unlines $ intersperse "" $ map show used) ++ "\n\n\n"
putStrLn $ "Finished calculating used"
Expand Down Expand Up @@ -136,7 +164,8 @@ codegenJs conf ci =
, doPartials (partialApplications stats), "\n"
, doHiddenClasses (hiddenClasses stats), "\n"
, out, "\n"
, jsName (sMN 0 "runMain"), "();\n"
, if iface then T.concat ["module.exports = {\n", T.intercalate ",\n" $ concatMap (makeExportDecls defs) (exportDecls ci), "\n};\n"]
else jsName (sMN 0 "runMain") `T.append` "();\n"
, "}.call(this))"
, footer conf
]
Expand Down Expand Up @@ -223,13 +252,17 @@ getConsId n =
Just (LConstructor _ conId arity) -> pure (conId, arity)
_ -> error $ "Internal JS Backend error " ++ showCG n ++ " is not a constructor."

getArgList' :: Name -> Map Name LDecl -> Maybe [Name]
getArgList' n defs =
case Map.lookup n defs of
Just (LFun _ _ a _) -> Just a
_ -> Nothing

getArgList :: Name -> State CGBodyState (Maybe [Name])
getArgList n =
do
st <- get
case Map.lookup n (defs st) of
Just (LFun _ _ a _) -> pure $ Just a
_ -> pure Nothing
pure $ getArgList' n (defs st)

data BodyResTarget = ReturnBT
| DecBT Text
Expand Down
4 changes: 3 additions & 1 deletion src/Idris/REPL.hs
Original file line number Diff line number Diff line change
Expand Up @@ -79,6 +79,7 @@ import Idris.REPL.Parser
import Idris.Termination
import Idris.TypeSearch (searchByType)
import Idris.WhoCalls
import IRTS.CodegenCommon
import IRTS.Compiler
import Network
import Prelude hiding (id, (.), (<$>))
Expand Down Expand Up @@ -1280,7 +1281,8 @@ process fn (Compile codegen f)
return (Just m')
ir <- compile codegen f m
i <- getIState
runIO $ generate codegen (fst (head (idris_imported i))) ir
let ir' = ir {interfaces = iface}
runIO $ generate codegen (fst (head (idris_imported i))) ir'
where fc = fileFC "main"
process fn (LogLvl i) = setLogLevel i
process fn (LogCategory cs) = setLogCats cs
Expand Down
3 changes: 2 additions & 1 deletion test/TestData.hs
Original file line number Diff line number Diff line change
Expand Up @@ -72,7 +72,7 @@ testFamiliesData = [
( 18, ANY ),
( 19, ANY ),
( 20, ANY ),
( 21, ANY )]),
( 21, C_CG )]),
("bignum", "Bignum",
[ ( 1, ANY ),
( 2, ANY )]),
Expand Down Expand Up @@ -129,6 +129,7 @@ testFamiliesData = [
, ( 8, C_CG )
, ( 9, C_CG )
, ( 10, NODE_CG )
, ( 11, NODE_CG )
]),
("folding", "Folding",
[ ( 1, ANY )]),
Expand Down
2 changes: 2 additions & 0 deletions test/ffi011/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,2 @@
Ready to show...
[10, 20, 30, 40]
7 changes: 7 additions & 0 deletions test/ffi011/ffi011
Original file line number Diff line number Diff line change
@@ -0,0 +1,7 @@
l = require("./lib.js");

x = l.cons(10, l.cons(20, l.nil()));
y = l.cons(30, l.cons(40, l.nil()));
z = l.addLists(x, y);

console.log(l.showList(z));
27 changes: 27 additions & 0 deletions test/ffi011/ffi011.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,27 @@
module Lib

import Data.List

addLists : List Int -> List Int -> List Int
addLists xs ys = xs ++ ys

nil : List Int
nil = []

cons : Int -> List Int -> List Int
cons x xs = x :: xs

show' : List Int -> JS_IO String
show' xs = do putStrLn' "Ready to show..."
pure (show xs)

testList : FFI_Export FFI_JS "" []
testList = Data (List Int) "ListInt" $
Data (List Nat) "ListNat" $
Fun addLists "addLists" $
Fun nil "nil" $
Fun cons "cons" $
Data Nat "Nat" $
Fun Strings.length "lengthS" $
Fun show' "showList" $
End
4 changes: 4 additions & 0 deletions test/ffi011/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,4 @@
#!/usr/bin/env bash
${IDRIS:-idris} $@ ffi011.idr --interface -o lib.js
node ./ffi011
rm -f *.ibc lib.js

0 comments on commit 6d954d8

Please sign in to comment.