Generate an executable via CC

This builds a .o from the generated C, and statically links with the
libidris2_support library. It doesn't yet dynamically link with any
additional libraries.
This commit is contained in:
Edwin Brady 2020-10-11 18:35:51 +01:00
parent 788bea3906
commit 3007b5a99d
3 changed files with 107 additions and 24 deletions

View File

@ -61,4 +61,5 @@ or via the `IDRIS2_CG` environment variable.
racket
gambit
javascript
refc
custom

View File

@ -0,0 +1,35 @@
*************************
C with Reference Counting
*************************
There is an experimental code generator which compiles to an executable via C,
using a reference counting garbage collector. This is intended as a lightweight
(i.e. minimal dependencies) code generator that can be ported to multiple
platforms, especially those with memory constraints.
Performance is not as good as the Scheme based code generators, partly because
the reference counting has not yet had any optimisation, and partly because of
the limitations of C. However, the main goal is portability: the generated
code should run on any platform that supports a C compiler.
This code generator can be accessed via the REPL command:
::
Main> :set cg refc
Alternatively, you can set it via the ``IDRIS2_CG`` environment variable:
::
$ export IDRIS2_CG=refc
The C compiler it invokes is determined by either the ``IDRIS2_CC`` or ``CC``
environment variables. If neither is set, it uses ``cc``.
This code generator does not yet support `:exec`, just `:c`.
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.

View File

@ -5,22 +5,35 @@ import Compiler.CompileExpr
import Compiler.LambdaLift
import Compiler.ANF
import Compiler.Inline
import Core.Context
import Core.Directory
import Core.Name
import Core.Options
import Core.TT
import Data.IORef
import Data.List
import Data.NameMap
import Data.Nat
import Data.Strings
import Data.Vect
import System
import System.Info
import System.File
import Utils.Hex
import Data.Nat
import Data.Strings
import Data.IORef
import Idris.Version
import Utils.Hex
import Utils.Path
findCC : IO String
findCC
= do Just cc <- getEnv "IDRIS2_CC"
| Nothing => do Just cc <- getEnv "CC"
| Nothing => pure "cc"
pure cc
pure cc
toString : List Char -> String
toString [] = ""
@ -1073,36 +1086,70 @@ footer = do
export
executeExpr : Ref Ctxt Defs -> (execDir : String) -> ClosedTerm -> Core ()
executeExpr c _ tm
= do coreLift $ system "false"
= do coreLift $ putStrLn "Execute expression not yet implemented for refc"
coreLift $ system "false"
pure ()
export
compileExpr : UsePhase
-> Ref Ctxt Defs
-> (tmpDir : String)
-> (execDir : String)
-> (outputDir : String)
-> ClosedTerm
-> (outfile : String)
-> Core (Maybe String)
compileExpr ANF c _ _ tm outfile = do
let outn = "build/code/" ++ outfile ++ ".c"
cdata <- getCompileData ANF tm
let defs = anf cdata
newRef ArgCounter 0
newRef FunctionDefinitions []
newRef TemporaryVariableTracker []
newRef OutfileText []
newRef ExternalLibs []
newRef IndentLevel 0
traverse (\(n, d) => createCFunctions n d) defs
header -- added after the definition traversal in order to add all encountered function defintions
footer
fileContent <- get OutfileText
let code = fastAppend (map (++ "\n") fileContent)
compileExpr ANF c _ outputDir tm outfile =
do let outn = outputDir </> outfile ++ ".c"
let outobj = outputDir </> outfile ++ ".o"
let outexec = outputDir </> outfile
coreLift (writeFile outn code)
coreLift $ putStrLn $ "File written to " ++ outn
pure (Just outn)
coreLift $ mkdirAll outputDir
cdata <- getCompileData ANF tm
let defs = anf cdata
newRef ArgCounter 0
newRef FunctionDefinitions []
newRef TemporaryVariableTracker []
newRef OutfileText []
newRef ExternalLibs []
newRef IndentLevel 0
traverse (\(n, d) => createCFunctions n d) defs
header -- added after the definition traversal in order to add all encountered function defintions
footer
fileContent <- get OutfileText
let code = fastAppend (map (++ "\n") fileContent)
coreLift (writeFile outn code)
coreLift $ putStrLn $ "Generated C file " ++ outn
cc <- coreLift findCC
dirs <- getDirs
let runccobj = cc ++ " -c " ++ outn ++ " -o " ++ outobj ++ " " ++
"-I" ++ fullprefix_dir dirs "refc " ++
"-I" ++ fullprefix_dir dirs "include"
let runcc = cc ++ " " ++ outobj ++ " -o " ++ outexec ++ " " ++
fullprefix_dir dirs "lib" </> "libidris2_support.a" ++ " " ++
"-lidris2_refc " ++
"-L" ++ fullprefix_dir dirs "refc " ++
clibdirs (lib_dirs dirs)
coreLift $ putStrLn runccobj
ok <- coreLift $ system runccobj
if ok == 0
then do coreLift $ putStrLn runcc
ok <- coreLift $ system runcc
if ok == 0
then pure (Just outexec)
else pure Nothing
else pure Nothing
where
fullprefix_dir : Dirs -> String -> String
fullprefix_dir dirs sub
= prefix_dir dirs </> "idris2-" ++ showVersion False version </> sub
clibdirs : List String -> String
clibdirs ds = concat (map (\d => "-L" ++ d ++ " ") ds)
compileExpr _ _ _ _ _ _ = pure Nothing
export