Allow overriding the build directory and the output directory

The output directory was previously called the executable directory, but I changed it because the output is not always an executable (depending on the code generator).

The code generators can now distinguish between where to place the (temporary) build files and the resulting output files.
This commit is contained in:
Christian Rasmussen 2020-06-16 22:04:38 +02:00
parent 8c556d0c26
commit 1f0ca85678
12 changed files with 110 additions and 67 deletions

View File

@ -27,13 +27,13 @@ Now create a file containing
import Compiler.Common
import Idris.Driver
compile : Ref Ctxt Defs -> (execDir : String) ->
compile : Ref Ctxt Defs -> (tmpDir : String) -> (outputDir : String) ->
ClosedTerm -> (outfile : String) -> Core (Maybe String)
compile defs dir term file = do coreLift $ putStrLn "I'd rather not."
compile defs tmpDir outputDir term file = do coreLift $ putStrLn "I'd rather not."
pure $ Nothing
execute : Ref Ctxt Defs -> (execDir : String) -> ClosedTerm -> Core ()
execute defs dir term = do coreLift $ putStrLn "Maybe in an hour."
execute : Ref Ctxt Defs -> (tmpDir : String) -> ClosedTerm -> Core ()
execute defs tmpDir term = do coreLift $ putStrLn "Maybe in an hour."
lazyCodegen : Codegen
lazyCodegen = MkCG compile execute

View File

@ -29,10 +29,10 @@ public export
record Codegen where
constructor MkCG
||| Compile an Idris 2 expression, saving it to a file.
compileExpr : Ref Ctxt Defs -> (execDir : String) ->
compileExpr : Ref Ctxt Defs -> (tmpDir : String) -> (outputDir : String) ->
ClosedTerm -> (outfile : String) -> Core (Maybe String)
||| Execute an Idris 2 expression directly.
executeExpr : Ref Ctxt Defs -> (execDir : String) -> ClosedTerm -> Core ()
executeExpr : Ref Ctxt Defs -> (tmpDir : String) -> ClosedTerm -> Core ()
-- Say which phase of compilation is the last one to use - it saves time if
-- you only ask for what you need.
@ -80,10 +80,13 @@ compile : {auto c : Ref Ctxt Defs} ->
Codegen ->
ClosedTerm -> (outfile : String) -> Core (Maybe String)
compile {c} cg tm out
= do makeExecDirectory
d <- getDirs
= do d <- getDirs
let tmpDir = execBuildDir d
let outputDir = outputDirWithDefault d
ensureDirectoryExists tmpDir
ensureDirectoryExists outputDir
logTime "Code generation overall" $
compileExpr cg c (exec_dir d) tm out
compileExpr cg c tmpDir outputDir tm out
||| execute
||| As with `compile`, produce a functon that executes
@ -92,9 +95,10 @@ export
execute : {auto c : Ref Ctxt Defs} ->
Codegen -> ClosedTerm -> Core ()
execute {c} cg tm
= do makeExecDirectory
d <- getDirs
executeExpr cg c (exec_dir d) tm
= do d <- getDirs
let tmpDir = execBuildDir d
ensureDirectoryExists tmpDir
executeExpr cg c tmpDir tm
pure ()
-- If an entry isn't already decoded, get the minimal entry we need for

View File

@ -432,22 +432,22 @@ makeShWindows chez outShRel appdir outAbs
pure ()
||| Chez Scheme implementation of the `compileExpr` interface.
compileExpr : Bool -> Ref Ctxt Defs -> (execDir : String) ->
compileExpr : Bool -> Ref Ctxt Defs -> (tmpDir : String) -> (outputDir : String) ->
ClosedTerm -> (outfile : String) -> Core (Maybe String)
compileExpr makeitso c execDir tm outfile
compileExpr makeitso c tmpDir outputDir tm outfile
= do let appDirRel = outfile ++ "_app" -- relative to build dir
let appDirGen = execDir </> appDirRel -- relative to here
let appDirGen = outputDir </> appDirRel -- relative to here
coreLift $ mkdirAll appDirGen
Just cwd <- coreLift currentDir
| Nothing => throw (InternalError "Can't get current directory")
let outSsFile = appDirRel </> outfile <.> "ss"
let outSoFile = appDirRel </> outfile <.> "so"
let outSsAbs = cwd </> execDir </> outSsFile
let outSoAbs = cwd </> execDir </> outSoFile
let outSsAbs = cwd </> outputDir </> outSsFile
let outSoAbs = cwd </> outputDir </> outSoFile
chez <- coreLift $ findChez
compileToSS c appDirGen tm outSsAbs
logTime "Make SO" $ when makeitso $ compileToSO chez appDirGen outSsAbs
let outShRel = execDir </> outfile
let outShRel = outputDir </> outfile
if isWindows
then makeShWindows chez outShRel appDirRel (if makeitso then outSoFile else outSsFile)
else makeSh outShRel appDirRel (if makeitso then outSoFile else outSsFile)
@ -456,9 +456,9 @@ compileExpr makeitso c execDir tm outfile
||| Chez Scheme implementation of the `executeExpr` interface.
||| This implementation simply runs the usual compiler, saving it to a temp file, then interpreting it.
executeExpr : Ref Ctxt Defs -> (execDir : String) -> ClosedTerm -> Core ()
executeExpr c execDir tm
= do Just sh <- compileExpr False c execDir tm "_tmpchez"
executeExpr : Ref Ctxt Defs -> (tmpDir : String) -> ClosedTerm -> Core ()
executeExpr c tmpDir tm
= do Just sh <- compileExpr False c tmpDir tmpDir tm "_tmpchez"
| Nothing => throw (InternalError "compileExpr returned Nothing")
coreLift $ system sh
pure ()

View File

@ -383,24 +383,25 @@ compileToSCM c tm outfile
| Left err => throw (FileErr outfile err)
pure $ mapMaybe fst fgndefs
compileExpr : Ref Ctxt Defs -> (execDir : String) ->
compileExpr : Ref Ctxt Defs -> (tmpDir : String) -> (outputDir : String) ->
ClosedTerm -> (outfile : String) -> Core (Maybe String)
compileExpr c execDir tm outfile
= do let outn = execDir </> outfile <.> "scm"
libsname <- compileToSCM c tm outn
compileExpr c tmpDir outputDir tm outfile
= do let srcPath = tmpDir </> outfile <.> "scm"
let execPath = outputDir </> outfile
libsname <- compileToSCM c tm srcPath
libsfile <- traverse findLibraryFile $ map (<.> "a") (nub libsname)
gsc <- coreLift findGSC
let cmd = gsc ++
" -exe -cc-options \"-Wno-implicit-function-declaration\" -ld-options \"" ++
(showSep " " libsfile) ++ "\" " ++ outn
(showSep " " libsfile) ++ "\" -o \"" ++ execPath ++ "\" \"" ++ srcPath ++ "\""
ok <- coreLift $ system cmd
if ok == 0
then pure (Just (execDir </> outfile))
then pure (Just execPath)
else pure Nothing
executeExpr : Ref Ctxt Defs -> (execDir : String) -> ClosedTerm -> Core ()
executeExpr c execDir tm
= do outn <- compileExpr c execDir tm "_tmpgambit"
executeExpr : Ref Ctxt Defs -> (tmpDir : String) -> ClosedTerm -> Core ()
executeExpr c tmpDir tm
= do outn <- compileExpr c tmpDir tmpDir tm "_tmpgambit"
case outn of
-- TODO: on windows, should add exe extension
Just outn => map (const ()) $ coreLift $ system outn

View File

@ -395,11 +395,11 @@ makeShWindows racket outShRel appdir outAbs
| Left err => throw (FileErr outShRel err)
pure ()
compileExpr : Bool -> Ref Ctxt Defs -> (execDir : String) ->
compileExpr : Bool -> Ref Ctxt Defs -> (tmpDir : String) -> (outputDir : String) ->
ClosedTerm -> (outfile : String) -> Core (Maybe String)
compileExpr mkexec c execDir tm outfile
compileExpr mkexec c tmpDir outputDir tm outfile
= do let appDirRel = outfile ++ "_app" -- relative to build dir
let appDirGen = execDir </> appDirRel -- relative to here
let appDirGen = outputDir </> appDirRel -- relative to here
coreLift $ mkdirAll appDirGen
Just cwd <- coreLift currentDir
| Nothing => throw (InternalError "Can't get current directory")
@ -407,8 +407,8 @@ compileExpr mkexec c execDir tm outfile
let ext = if isWindows then ".exe" else ""
let outRktFile = appDirRel </> outfile <.> "rkt"
let outBinFile = appDirRel </> outfile <.> ext
let outRktAbs = cwd </> execDir </> outRktFile
let outBinAbs = cwd </> execDir </> outBinFile
let outRktAbs = cwd </> outputDir </> outRktFile
let outBinAbs = cwd </> outputDir </> outBinFile
compileToRKT c appDirGen tm outRktAbs
raco <- coreLift findRacoExe
@ -421,7 +421,7 @@ compileExpr mkexec c execDir tm outfile
else pure 0
if ok == 0
then do -- TODO: add launcher script
let outShRel = execDir </> outfile
let outShRel = outputDir </> outfile
the (Core ()) $ if isWindows
then if mkexec
then makeShWindows "" outShRel appDirRel outBinFile
@ -433,9 +433,9 @@ compileExpr mkexec c execDir tm outfile
pure (Just outShRel)
else pure Nothing
executeExpr : Ref Ctxt Defs -> (execDir : String) -> ClosedTerm -> Core ()
executeExpr c execDir tm
= do Just sh <- compileExpr False c execDir tm "_tmpracket"
executeExpr : Ref Ctxt Defs -> (tmpDir : String) -> ClosedTerm -> Core ()
executeExpr c tmpDir tm
= do Just sh <- compileExpr False c tmpDir tmpDir tm "_tmpracket"
| Nothing => throw (InternalError "compileExpr returned Nothing")
coreLift $ system sh
pure ()

View File

@ -1932,10 +1932,10 @@ setBuildDir dir
put Ctxt (record { options->dirs->build_dir = dir } defs)
export
setExecDir : {auto c : Ref Ctxt Defs} -> String -> Core ()
setExecDir dir
setOutputDir : {auto c : Ref Ctxt Defs} -> Maybe String -> Core ()
setOutputDir dir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->exec_dir = dir } defs)
put Ctxt (record { options->dirs->output_dir = dir } defs)
export
setSourceDir : {auto c : Ref Ctxt Defs} -> Maybe String -> Core ()

View File

@ -136,12 +136,10 @@ makeBuildDirectory ns
export
covering
makeExecDirectory : {auto c : Ref Ctxt Defs} ->
Core ()
makeExecDirectory
= do d <- getDirs
Right _ <- coreLift $ mkdirAll (exec_dir d)
| Left err => throw (FileErr (exec_dir d) err)
ensureDirectoryExists : String -> Core ()
ensureDirectoryExists dir
= do Right _ <- coreLift $ mkdirAll dir
| Left err => throw (FileErr dir err)
pure ()
-- Given a source file, return the name of the ttc file to generate

View File

@ -7,6 +7,7 @@ import Utils.Binary
import Utils.Path
import Data.List
import Data.Maybe
import Data.Strings
import System.Info
@ -19,23 +20,31 @@ record Dirs where
working_dir : String
source_dir : Maybe String -- source directory, relative to working directory
build_dir : String -- build directory, relative to working directory
exec_dir : String -- executable directory, relative to working directory
output_dir : Maybe String -- output directory, relative to working directory
dir_prefix : String -- installation prefix, for finding data files (e.g. run time support)
extra_dirs : List String -- places to look for import files
lib_dirs : List String -- places to look for libraries (for code generation)
data_dirs : List String -- places to look for data file
export
execBuildDir : Dirs -> String
execBuildDir d = build_dir d </> "exec"
export
outputDirWithDefault : Dirs -> String
outputDirWithDefault d = fromMaybe (build_dir d </> "exec") (output_dir d)
public export
toString : Dirs -> String
toString (MkDirs wdir sdir bdir edir dfix edirs ldirs ddirs) =
unlines [ "+ Working Directory :: " ++ show wdir
, "+ Source Directory :: " ++ show sdir
, "+ Build Directory :: " ++ show bdir
, "+ Executable Directory :: " ++ show edir
, "+ Installation Prefix :: " ++ show dfix
, "+ Extra Directories :: " ++ show edirs
toString d@(MkDirs wdir sdir bdir odir dfix edirs ldirs ddirs) =
unlines [ "+ Working Directory :: " ++ show wdir
, "+ Source Directory :: " ++ show sdir
, "+ Build Directory :: " ++ show bdir
, "+ Output Directory :: " ++ show (outputDirWithDefault d)
, "+ Installation Prefix :: " ++ show dfix
, "+ Extra Directories :: " ++ show edirs
, "+ CG Library Directories :: " ++ show ldirs
, "+ Data Directories :: " ++ show ddirs]
, "+ Data Directories :: " ++ show ddirs]
public export
data CG = Chez
@ -149,8 +158,7 @@ getCG : Options -> String -> Maybe CG
getCG o cg = lookup (toLower cg) (availableCGs o)
defaultDirs : Dirs
defaultDirs = MkDirs "." Nothing "build"
("build" </> "exec")
defaultDirs = MkDirs "." Nothing "build" Nothing
"/usr/local" ["."] [] []
defaultPPrint : PPrinter

View File

@ -50,6 +50,10 @@ data CLOpt
SetCG String |
||| Don't implicitly import Prelude
NoPrelude |
||| Set build directory
BuildDir String |
||| Set output directory
OutputDir String |
||| Show the installation prefix
ShowPrefix |
||| Display Idris version
@ -149,6 +153,10 @@ options = [MkOpt ["--check", "-c"] [] [CheckOnly]
(Just $ "Set code generator " ++ showDefault (codegen defaultSession)),
MkOpt ["--package", "-p"] [Required "package"] (\f => [PkgPath f])
(Just "Add a package as a dependency"),
MkOpt ["--build-dir"] [Required "dir"] (\d => [BuildDir d])
(Just $ "Set build directory"),
MkOpt ["--output-dir"] [Required "dir"] (\d => [OutputDir d])
(Just $ "Set output directory"),
optSeparator,
MkOpt ["--prefix"] [] [ShowPrefix]

View File

@ -57,6 +57,8 @@ record PkgDesc where
executable : Maybe String -- name of executable
options : Maybe (FC, String)
sourcedir : Maybe String
builddir : Maybe String
outputdir : Maybe String
prebuild : Maybe (FC, String) -- Script to run before building
postbuild : Maybe (FC, String) -- Script to run after building
preinstall : Maybe (FC, String) -- Script to run after building, before installing
@ -81,6 +83,8 @@ Show PkgDesc where
maybe "" (\m => "Exec: " ++ m ++ "\n") (executable pkg) ++
maybe "" (\m => "Opts: " ++ snd m ++ "\n") (options pkg) ++
maybe "" (\m => "SourceDir: " ++ m ++ "\n") (sourcedir pkg) ++
maybe "" (\m => "BuildDir: " ++ m ++ "\n") (builddir pkg) ++
maybe "" (\m => "OutputDir: " ++ m ++ "\n") (outputdir pkg) ++
maybe "" (\m => "Prebuild: " ++ snd m ++ "\n") (prebuild pkg) ++
maybe "" (\m => "Postbuild: " ++ snd m ++ "\n") (postbuild pkg) ++
maybe "" (\m => "Preinstall: " ++ snd m ++ "\n") (preinstall pkg) ++
@ -94,7 +98,7 @@ initPkgDesc pname
Nothing Nothing Nothing Nothing Nothing
[] []
Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing
Nothing Nothing
Nothing Nothing Nothing Nothing
data DescField : Type where
PVersion : FC -> String -> DescField
@ -112,6 +116,8 @@ data DescField : Type where
PExec : String -> DescField
POpts : FC -> String -> DescField
PSourceDir : FC -> String -> DescField
PBuildDir : FC -> String -> DescField
POutputDir : FC -> String -> DescField
PPrebuild : FC -> String -> DescField
PPostbuild : FC -> String -> DescField
PPreinstall : FC -> String -> DescField
@ -133,6 +139,8 @@ field fname
<|> strField POpts "options"
<|> strField POpts "opts"
<|> strField PSourceDir "sourcedir"
<|> strField PBuildDir "builddir"
<|> strField POutputDir "outputdir"
<|> strField PPrebuild "prebuild"
<|> strField PPostbuild "postbuild"
<|> strField PPreinstall "preinstall"
@ -204,6 +212,8 @@ addField (PMainMod loc n) pkg = do put MainMod (Just (loc, n))
addField (PExec e) pkg = pure $ record { executable = Just e } pkg
addField (POpts fc e) pkg = pure $ record { options = Just (fc, e) } pkg
addField (PSourceDir fc a) pkg = pure $ record { sourcedir = Just a } pkg
addField (PBuildDir fc a) pkg = pure $ record { builddir = Just a } pkg
addField (POutputDir fc a) pkg = pure $ record { outputdir = Just a } pkg
addField (PPrebuild fc e) pkg = pure $ record { prebuild = Just (fc, e) } pkg
addField (PPostbuild fc e) pkg = pure $ record { postbuild = Just (fc, e) } pkg
addField (PPreinstall fc e) pkg = pure $ record { preinstall = Just (fc, e) } pkg
@ -391,8 +401,6 @@ clean : {auto c : Ref Ctxt Defs} ->
Core ()
clean pkg opts -- `opts` is not used but might be in the future
= do defs <- get Ctxt
let build = build_dir (dirs (options defs))
let exec = exec_dir (dirs (options defs))
runScript (preclean pkg)
let pkgmods = maybe
(map fst (modules pkg))
@ -404,8 +412,9 @@ clean pkg opts -- `opts` is not used but might be in the future
(x :: xs) => Just (xs, x)) pkgmods
Just srcdir <- coreLift currentDir
| Nothing => throw (InternalError "Can't get current directory")
let builddir = srcdir </> build </> "ttc"
let execdir = srcdir </> exec
let d = dirs (options defs)
let builddir = srcdir </> build_dir d </> "ttc"
let outputdir = srcdir </> outputDirWithDefault d
-- the usual pair syntax breaks with `No such variable a` here for some reason
let pkgTrie = the (StringTrie (List String)) $
foldl (\trie, ksv =>
@ -417,7 +426,7 @@ clean pkg opts -- `opts` is not used but might be in the future
(\ks => map concat . traverse (deleteBin builddir ks))
pkgTrie
deleteFolder builddir []
maybe (pure ()) (\e => delete (execdir </> e))
maybe (pure ()) (\e => delete (outputdir </> e))
(executable pkg)
runScript (postclean pkg)
where
@ -473,6 +482,8 @@ processPackage cmd file opts
| Left (FileFail err) => throw (FileErr file err)
| Left err => throw (ParseFail (getParseErrorLoc file err) err)
pkg <- addFields fs (initPkgDesc pname)
maybe (pure ()) setBuildDir (builddir pkg)
setOutputDir (outputdir pkg)
case cmd of
Build => do [] <- build pkg opts
| errs => coreLift (exitWith (ExitFailure 1))
@ -504,6 +515,8 @@ errorMsg = unlines
, " --dumpvmcode <file>"
, " --debug-elab-check"
, " --codegen <cg>"
, " --build-dir <dir>"
, " --output-dir <dir>"
]
@ -511,7 +524,6 @@ filterPackageOpts : POptsFilterResult -> List CLOpt -> Core (POptsFilterResult)
filterPackageOpts acc Nil = pure acc
filterPackageOpts acc (Package cmd f ::xs) = filterPackageOpts (record {pkgDetails = Just (cmd, f)} acc) xs
filterPackageOpts acc (SetCG f ::xs) = filterPackageOpts (record {oopts $= (SetCG f::)} acc) xs
filterPackageOpts acc (Quiet ::xs) = filterPackageOpts (record {oopts $= (Quiet::)} acc) xs
filterPackageOpts acc (Verbose ::xs) = filterPackageOpts (record {oopts $= (Verbose::)} acc) xs
filterPackageOpts acc (Timing ::xs) = filterPackageOpts (record {oopts $= (Timing::)} acc) xs
@ -519,6 +531,9 @@ filterPackageOpts acc (DumpCases f ::xs) = filterPackageOpts (record {oopts $=
filterPackageOpts acc (DumpLifted f ::xs) = filterPackageOpts (record {oopts $= (DumpLifted f::)} acc) xs
filterPackageOpts acc (DumpVMCode f ::xs) = filterPackageOpts (record {oopts $= (DumpVMCode f::)} acc) xs
filterPackageOpts acc (DebugElabCheck::xs) = filterPackageOpts (record {oopts $= (DebugElabCheck::)} acc) xs
filterPackageOpts acc (SetCG f ::xs) = filterPackageOpts (record {oopts $= (SetCG f::)} acc) xs
filterPackageOpts acc (BuildDir f ::xs) = filterPackageOpts (record {oopts $= (BuildDir f::)} acc) xs
filterPackageOpts acc (OutputDir f ::xs) = filterPackageOpts (record {oopts $= (OutputDir f::)} acc) xs
filterPackageOpts acc (x::xs) = pure (record {hasError = True} acc)
@ -564,7 +579,8 @@ findIpkg fname
| Left (FileFail err) => throw (FileErr ipkgn err)
| Left err => throw (ParseFail (getParseErrorLoc ipkgn err) err)
pkg <- addFields fs (initPkgDesc pname)
setSourceDir (sourcedir pkg)
maybe (pure ()) setBuildDir (builddir pkg)
setOutputDir (outputdir pkg)
processOptions (options pkg)
loadDependencies (depends pkg)
case fname of

View File

@ -74,6 +74,12 @@ preOptions (SetCG e :: opts)
preOptions (PkgPath p :: opts)
= do addPkgDir p
preOptions opts
preOptions (BuildDir d :: opts)
= do setBuildDir d
preOptions opts
preOptions (OutputDir d :: opts)
= do setOutputDir (Just d)
preOptions opts
preOptions (Directory d :: opts)
= do defs <- get Ctxt
dirOption (dirs (options defs)) d

View File

@ -10,6 +10,8 @@ Overridable options are:
--dumpvmcode <file>
--debug-elab-check
--codegen <cg>
--build-dir <dir>
--output-dir <dir>
Packages must have an '.ipkg' extension: "malformed-package-name".
Uncaught error: File error (non-existent-package.ipkg): File Not Found