mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-27 13:40:15 +03:00
add extraRuntime
option for Scheme backends (#578)
This commit is contained in:
parent
da78ac4783
commit
3b49b10832
@ -6,6 +6,7 @@ Alex Silva
|
||||
Andre Kuhlenschmidt
|
||||
André Videla
|
||||
Andy Lok
|
||||
Anthony Lodi
|
||||
Arnaud Bailly
|
||||
Brian Wignall
|
||||
Christian Rasmussen
|
||||
@ -22,7 +23,6 @@ Johann Rudloff
|
||||
Kamil Shakirov
|
||||
Bryn Keller
|
||||
Kevin Boulain
|
||||
lodi
|
||||
LuoChen
|
||||
Marc Petit-Huguenin
|
||||
MarcelineVQ
|
||||
|
@ -45,3 +45,29 @@ that.
|
||||
Chez Scheme is the default code generator, so if you invoke ``idris2`` with the
|
||||
``-o execname`` flag, it will generate an executable script
|
||||
``build/exec/execname``, again with support files in ``build/exec/execname_app``.
|
||||
|
||||
|
||||
Chez Directives
|
||||
===============
|
||||
|
||||
* ``--directive extraRuntime=<path>``
|
||||
|
||||
Embed Scheme source from ``<path>`` directly into generated output. Can be specified more than
|
||||
once, in which case all given files will be included in the order specified.
|
||||
|
||||
.. code-block:: scheme
|
||||
|
||||
; extensions.scm
|
||||
(define (my-mul a b)
|
||||
(* a b))
|
||||
|
||||
|
||||
.. code-block:: idris
|
||||
|
||||
-- Main.idr
|
||||
%foreign "scheme:my-mul"
|
||||
myMul : Int -> Int -> Int
|
||||
|
||||
.. code-block::
|
||||
|
||||
$ idris2 --codegen chez --directive extraRuntime=/path/to/extensions.scm -o main Main.idr
|
||||
|
@ -40,3 +40,29 @@ You can also execute an expression directly:
|
||||
|
||||
Again, ``expr`` must have type ``IO ()``. This will generate a temporary
|
||||
Scheme file, and execute the Gambit interpreter on it.
|
||||
|
||||
|
||||
Gambit Directives
|
||||
=================
|
||||
|
||||
* ``--directive extraRuntime=<path>``
|
||||
|
||||
Embed Scheme source from ``<path>`` directly into generated output. Can be specified more than
|
||||
once, in which case all given files will be included in the order specified.
|
||||
|
||||
.. code-block:: scheme
|
||||
|
||||
; extensions.scm
|
||||
(define (my-mul a b)
|
||||
(* a b))
|
||||
|
||||
|
||||
.. code-block:: idris
|
||||
|
||||
-- Main.idr
|
||||
%foreign "scheme:my-mul"
|
||||
myMul : Int -> Int -> Int
|
||||
|
||||
.. code-block::
|
||||
|
||||
$ idris2 --codegen chez --directive extraRuntime=/path/to/extensions.scm -o main Main.idr
|
||||
|
@ -43,3 +43,29 @@ Again, ``expr`` must have type ``IO ()``. This will generate a temporary
|
||||
executable script ``_tmpracket`` in the ``build/exec`` directory, and execute
|
||||
that, without compiling to a binary first (so the resulting Racket code is
|
||||
interpreted).
|
||||
|
||||
|
||||
Racket Directives
|
||||
=================
|
||||
|
||||
* ``--directive extraRuntime=<path>``
|
||||
|
||||
Embed Scheme source from ``<path>`` directly into generated output. Can be specified more than
|
||||
once, in which case all given files will be included in the order specified.
|
||||
|
||||
.. code-block:: scheme
|
||||
|
||||
; extensions.scm
|
||||
(define (my-mul a b)
|
||||
(* a b))
|
||||
|
||||
|
||||
.. code-block:: idris
|
||||
|
||||
-- Main.idr
|
||||
%foreign "scheme:my-mul"
|
||||
myMul : Int -> Int -> Int
|
||||
|
||||
.. code-block::
|
||||
|
||||
$ idris2 --codegen chez --directive extraRuntime=/path/to/extensions.scm -o main Main.idr
|
||||
|
@ -407,3 +407,29 @@ copyLib (lib, fullname)
|
||||
Right _ <- coreLift $ writeToFile lib bin
|
||||
| Left err => throw (FileErr lib err)
|
||||
pure ()
|
||||
|
||||
|
||||
-- parses `--directive extraRuntime=/path/to/defs.scm` options for textual inclusion in generated
|
||||
-- source. Use with `%foreign "scheme:..."` declarations to write runtime-specific scheme calls.
|
||||
export
|
||||
getExtraRuntime : List String -> Core String
|
||||
getExtraRuntime directives
|
||||
= do fileContents <- traverse readPath paths
|
||||
pure $ concat $ intersperse "\n" fileContents
|
||||
where
|
||||
getArg : String -> Maybe String
|
||||
getArg directive =
|
||||
let (k,v) = break (== '=') directive
|
||||
in
|
||||
if (trim k) == "extraRuntime"
|
||||
then Just $ trim $ substr 1 (length v) v
|
||||
else Nothing
|
||||
|
||||
paths : List String
|
||||
paths = nub $ mapMaybe getArg $ reverse directives
|
||||
|
||||
readPath : String -> Core String
|
||||
readPath p = do
|
||||
Right contents <- coreLift $ readFile p
|
||||
| Left err => throw (FileErr p err)
|
||||
pure contents
|
||||
|
@ -385,8 +385,9 @@ compileToSS c appdir tm outfile
|
||||
main <- schExp chezExtPrim chezString 0 ctm
|
||||
chez <- coreLift findChez
|
||||
support <- readDataFile "chez/support.ss"
|
||||
extraRuntime <- getExtraRuntime ds
|
||||
let scm = schHeader chez (map snd libs) ++
|
||||
support ++ code ++
|
||||
support ++ extraRuntime ++ code ++
|
||||
concat (map fst fgndefs) ++
|
||||
"(collect-request-handler (lambda () (collect) (blodwen-run-finalisers)))\n" ++
|
||||
main ++ schFooter
|
||||
|
@ -371,8 +371,10 @@ compileToSCM c tm outfile
|
||||
let code = fastAppend (map snd fgndefs ++ compdefs)
|
||||
main <- schExp gambitPrim gambitString 0 ctm
|
||||
support <- readDataFile "gambit/support.scm"
|
||||
ds <- getDirectives Gambit
|
||||
extraRuntime <- getExtraRuntime ds
|
||||
foreign <- readDataFile "gambit/foreign.scm"
|
||||
let scm = showSep "\n" [schHeader, support, foreign, code, main]
|
||||
let scm = showSep "\n" [schHeader, support, extraRuntime, foreign, code, main]
|
||||
Right () <- coreLift $ writeFile outfile scm
|
||||
| Left err => throw (FileErr outfile err)
|
||||
pure $ mapMaybe fst fgndefs
|
||||
|
@ -372,8 +372,10 @@ compileToRKT c appdir tm outfile
|
||||
let code = fastAppend (map snd fgndefs ++ compdefs)
|
||||
main <- schExp racketPrim racketString 0 ctm
|
||||
support <- readDataFile "racket/support.rkt"
|
||||
ds <- getDirectives Racket
|
||||
extraRuntime <- getExtraRuntime ds
|
||||
let scm = schHeader (concat (map fst fgndefs)) ++
|
||||
support ++ code ++
|
||||
support ++ extraRuntime ++ code ++
|
||||
"(void " ++ main ++ ")\n" ++
|
||||
schFooter
|
||||
Right () <- coreLift $ writeFile outfile scm
|
||||
|
@ -125,7 +125,7 @@ chezTests
|
||||
"chez007", "chez008", "chez009", "chez010", "chez011", "chez012",
|
||||
"chez013", "chez014", "chez015", "chez016", "chez017", "chez018",
|
||||
"chez019", "chez020", "chez021", "chez022", "chez023", "chez024",
|
||||
"chez025", "chez026", "chez027", "chez028", "chez029",
|
||||
"chez025", "chez026", "chez027", "chez028", "chez029", "chez030",
|
||||
"reg001"]
|
||||
|
||||
nodeTests : List String
|
||||
|
6
tests/chez/chez030/ExtraRuntime.idr
Normal file
6
tests/chez/chez030/ExtraRuntime.idr
Normal file
@ -0,0 +1,6 @@
|
||||
%foreign "scheme:my-mul"
|
||||
myMul : Int -> Int -> Int
|
||||
|
||||
main : IO ()
|
||||
main = do
|
||||
putStrLn $ "6 * 7 = " ++ show (myMul 6 7)
|
1
tests/chez/chez030/expected
Normal file
1
tests/chez/chez030/expected
Normal file
@ -0,0 +1 @@
|
||||
6 * 7 = 42
|
2
tests/chez/chez030/extensions.scm
Normal file
2
tests/chez/chez030/extensions.scm
Normal file
@ -0,0 +1,2 @@
|
||||
(define (my-mul a b)
|
||||
(* a b))
|
4
tests/chez/chez030/run
Normal file
4
tests/chez/chez030/run
Normal file
@ -0,0 +1,4 @@
|
||||
$1 --no-banner --no-color --console-width 0 --directive extraRuntime=extensions.scm -o chez030 ExtraRuntime.idr
|
||||
./build/exec/chez030
|
||||
|
||||
rm -rf build
|
Loading…
Reference in New Issue
Block a user