Add ability to extend RefC backend to create further backends

This commit is contained in:
Robert Wright 2021-06-17 17:23:29 +01:00 committed by G. Allais
parent 657699a866
commit a8264f8f05
6 changed files with 111 additions and 10 deletions

View File

@ -33,3 +33,71 @@ Also note that, if you link with any dynamic libraries for interfacing with
C, you will need to arrange for them to be accessible via ``LD_LIBRARY_PATH``
when running the executable. The default Idris 2 support libraries are
statically linked.
Extending RefC
==============
RefC can be extended to produce a new backend for languages that support C
foreign functions. For example, a
`Python backend for Idris <https://github.com/madman-bob/idris2-python>`_.
In your backend, use the ``Compiler.RefC`` functions ``generateCSourceFile``,
``compileCObjectFile {asLibrary = True}``, and
``compileCFile {asShared = True}`` to generate a ``.so`` shared object file.
.. code-block:: idris
_ <- generateCSourceFile defs cSourceFile
_ <- compileCObjectFile {asLibrary = True} cSourceFile cObjectFile
_ <- compileCFile {asShared = True} cObjectFile cSharedObjectFile
To run a compiled Idris program, call the ``int main(int argc, char *argv[])``
function in the compiled ``.so`` file, with the arguments you wish to pass to
the running program.
For example, in Python:
.. code-block:: python
import ctypes
import sys
argc = len(sys.argv)
argv = (ctypes.c_char_p * argc)(*map(str.encode, sys.argv))
cdll = ctypes.CDLL("main.so")
cdll.main(argc, argv)
Extending RefC FFIs
-------------------
To make the generated C code recognize additional FFI languages beyond the
standard RefC FFIs, pass the ``additionalFFILangs`` option to
``generateCSourceFile``, with a list of the language identifiers your backend
recognizes.
.. code-block:: idris
_ <- generateCSourceFile {additionalFFILangs = ["python"]} defs cSourceFile
This will generate stub FFI function pointers in the generated C file, which
your backend should set to the appropriate C functions before ``main`` is
called.
Each ``%foreign "lang: funcName, opts"`` definition will produce a stub whose
name is given by ``cName (UN $ lang ++ "_" ++ funcName)``, of the appropriate
function pointer type.
So the ``%foreign`` function
.. code-block:: idris
%foreign "python: abs"
abs : Int -> Int
produces a stub ``python_abs``, which can be backpatched in Python by:
.. code-block:: python
abs_ptr = ctypes.CFUNCTYPE(ctypes.c_int64, ctypes.c_int64)(abs)
ctypes.c_void_p.in_dll(cdll, "python_abs").value = ctypes.cast(abs_ptr, ctypes.c_void_p).value

View File

@ -23,6 +23,7 @@ modules =
Compiler.ES.Node,
Compiler.ES.TailRec,
Compiler.RefC,
Compiler.RefC.CC,
Compiler.RefC.RefC,

4
src/Compiler/RefC.idr Normal file
View File

@ -0,0 +1,4 @@
module Compiler.RefC
import public Compiler.RefC.CC
import public Compiler.RefC.RefC

View File

@ -62,6 +62,7 @@ showcCleanString (c ::cs) = (showcCleanStringChar c) . showcCleanString cs
cCleanString : String -> String
cCleanString cs = showcCleanString (unpack cs) ""
export
cName : Name -> String
cName (NS ns n) = cCleanString (showNSWithSep "_" ns) ++ "_" ++ cName n
cName (UN n) = cCleanString n
@ -837,6 +838,13 @@ discardLastArgument : List ty -> List ty
discardLastArgument [] = []
discardLastArgument xs@(_ :: _) = init xs
additionalFFIStub : Name -> List CFType -> CFType -> String
additionalFFIStub name argTypes (CFIORes retType) = additionalFFIStub name (discardLastArgument argTypes) retType
additionalFFIStub name argTypes retType =
cTypeOfCFType retType ++
" (*" ++ cName name ++ ")(" ++
(concat $ intersperse ", " $ map cTypeOfCFType argTypes) ++ ") = (void*)missing_ffi;\n"
createCFunctions : {auto c : Ref Ctxt Defs}
-> {auto a : Ref ArgCounter Nat}
-> {auto f : Ref FunctionDefinitions (List String)}
@ -844,6 +852,7 @@ createCFunctions : {auto c : Ref Ctxt Defs}
-> {auto oft : Ref OutfileText Output}
-> {auto il : Ref IndentLevel Nat}
-> {auto h : Ref HeaderFiles (SortedSet String)}
-> {default [] additionalFFILangs : List String}
-> Name
-> ANFDef
-> Core ()
@ -887,11 +896,17 @@ createCFunctions n (MkACon tag arity nt) = do
createCFunctions n (MkAForeign ccs fargs ret) = do
case parseCC ["RefC", "C"] ccs of
Just (_, fctName :: extLibOpts) => do
case extLibOpts of
[lib, header] => addHeader header
_ => pure ()
case parseCC (additionalFFILangs ++ ["RefC", "C"]) ccs of
Just (lang, fctForeignName :: extLibOpts) => do
let isStandardFFI = Prelude.elem lang ["RefC", "C"]
let fctName = if isStandardFFI
then UN fctForeignName
else UN $ lang ++ "_" ++ fctForeignName
if isStandardFFI
then case extLibOpts of
[lib, header] => addHeader header
_ => pure ()
else emit EmptyFC $ additionalFFIStub fctName fargs ret
otherDefs <- get FunctionDefinitions
let fnDef = "Value *" ++ (cName n) ++ "(" ++ showSep ", " (replicate (length fargs) "Value *") ++ ");"
fn_arglist <- functionDefSignatureArglist n
@ -917,22 +932,22 @@ createCFunctions n (MkAForeign ccs fargs ret) = do
emitFDef n typeVarNameArgList
emit EmptyFC "{"
increaseIndentation
emit EmptyFC $ " // ffi call to " ++ fctName
emit EmptyFC $ " // ffi call to " ++ cName fctName
case ret of
CFIORes CFUnit => do
emit EmptyFC $ fctName
emit EmptyFC $ cName fctName
++ "("
++ showSep ", " (map (\(_, vn, vt) => extractValue vt vn) (discardLastArgument typeVarNameArgList))
++ ");"
emit EmptyFC "return NULL;"
CFIORes ret => do
emit EmptyFC $ cTypeOfCFType ret ++ " retVal = " ++ fctName
emit EmptyFC $ cTypeOfCFType ret ++ " retVal = " ++ cName fctName
++ "("
++ showSep ", " (map (\(_, vn, vt) => extractValue vt vn) (discardLastArgument typeVarNameArgList))
++ ");"
emit EmptyFC $ "return (Value*)" ++ packCFType ret "retVal" ++ ";"
_ => do
emit EmptyFC $ cTypeOfCFType ret ++ " retVal = " ++ fctName
emit EmptyFC $ cTypeOfCFType ret ++ " retVal = " ++ cName fctName
++ "("
++ showSep ", " (map (\(_, vn, vt) => extractValue vt vn) typeVarNameArgList)
++ ");"
@ -985,6 +1000,7 @@ executeExpr c _ tm
export
generateCSourceFile : {auto c : Ref Ctxt Defs}
-> {default [] additionalFFILangs : List String}
-> List (Name, ANFDef)
-> (outn : String)
-> Core ()
@ -995,7 +1011,7 @@ generateCSourceFile defs outn =
_ <- newRef OutfileText DList.Nil
_ <- newRef HeaderFiles empty
_ <- newRef IndentLevel 0
traverse_ (uncurry createCFunctions) defs
traverse_ (uncurry $ createCFunctions {additionalFFILangs}) defs
header -- added after the definition traversal in order to add all encountered function defintions
footer
fileContent <- get OutfileText

View File

@ -1,5 +1,15 @@
#include "runtime.h"
void missing_ffi()
{
fprintf(
stderr,
"Foreign function declared, but not defined.\n"
"Cannot call missing FFI - aborting.\n"
);
exit(1);
}
void push_Arglist(Value_Arglist *arglist, Value *arg)
{
if (arglist->filled >= arglist->total)

View File

@ -3,6 +3,8 @@
#include "cBackend.h"
void missing_ffi();
Value *apply_closure(Value *, Value *arg);
void push_Arglist(Value_Arglist *arglist, Value *arg);