Skip to content

Commit

Permalink
feat: lake: build without leanc
Browse files Browse the repository at this point in the history
  • Loading branch information
tydeu committed Nov 22, 2024
1 parent 4600bb1 commit d4a70f3
Show file tree
Hide file tree
Showing 6 changed files with 82 additions and 37 deletions.
37 changes: 27 additions & 10 deletions src/lake/Lake/Build/Common.lean
Original file line number Diff line number Diff line change
Expand Up @@ -306,7 +306,8 @@ which will be computed in the resulting `BuildJob` before building.
let extraDepTrace :=
return (← getLeanTrace).mix <| (pureHash traceArgs).mix platformTrace
buildFileAfterDep oFile srcJob (extraDepTrace := extraDepTrace) fun srcFile => do
compileO oFile srcFile (weakArgs ++ traceArgs) (← getLeanc)
let lean ← getLeanInstall
compileO oFile srcFile (lean.ccFlags ++ weakArgs ++ traceArgs) lean.cc

/-- Build a static library from object file jobs using the `ar` packaged with Lean. -/
def buildStaticLib
Expand All @@ -315,27 +316,42 @@ def buildStaticLib
buildFileAfterDepArray libFile oFileJobs fun oFiles => do
compileStaticLib libFile oFiles (← getLeanAr)

/-- Build a shared library by linking the results of `linkJobs` using `leanc`. -/
/--
Build a shared library by linking the results of `linkJobs`
using the Lean toolchain's C compiler.
-/
def buildLeanSharedLib
(libFile : FilePath) (linkJobs : Array (BuildJob FilePath))
(weakArgs traceArgs : Array String := #[])
: SpawnM (BuildJob FilePath) :=
let extraDepTrace :=
return (← getLeanTrace).mix <| (pureHash traceArgs).mix platformTrace
buildFileAfterDepArray libFile linkJobs (extraDepTrace := extraDepTrace) fun links => do
compileSharedLib libFile (links.map toString ++ weakArgs ++ traceArgs) (← getLeanc)
let lean ← getLeanInstall
let args := links.map toString ++ weakArgs ++ traceArgs ++ lean.ccLinkSharedFlags
compileSharedLib libFile args lean.cc


/-- Build an executable by linking the results of `linkJobs` using `leanc`. -/
/--
Build an executable by linking the results of `linkJobs`
using the Lean toolchain's linker.
-/
def buildLeanExe
(exeFile : FilePath) (linkJobs : Array (BuildJob FilePath))
(weakArgs traceArgs : Array String := #[])
(weakArgs traceArgs : Array String := #[]) (sharedLean : Bool := false)
: SpawnM (BuildJob FilePath) :=
let extraDepTrace :=
return (← getLeanTrace).mix <| (pureHash traceArgs).mix platformTrace
return (← getLeanTrace).mix <|
(pureHash sharedLean).mix <| (pureHash traceArgs).mix platformTrace
buildFileAfterDepArray exeFile linkJobs (extraDepTrace := extraDepTrace) fun links => do
compileExe exeFile links (weakArgs ++ traceArgs) (← getLeanc)
let lean ← getLeanInstall
let args := weakArgs ++ traceArgs ++ lean.ccLinkFlags sharedLean
compileExe exeFile links args lean.cc

/-- Build a shared library from a static library using `leanc`. -/
/--
Build a shared library from a static library using `leanc`
using the Lean toolchain's linker.
-/
def buildLeanSharedLibOfStatic
(staticLibJob : BuildJob FilePath)
(weakArgs traceArgs : Array String := #[])
Expand All @@ -347,11 +363,12 @@ def buildLeanSharedLibOfStatic
#[s!"-Wl,-force_load,{staticLib}"]
else
#["-Wl,--whole-archive", staticLib.toString, "-Wl,--no-whole-archive"]
let lean ← getLeanInstall
let depTrace := staticTrace.mix <|
(← getLeanTrace).mix <| (pureHash traceArgs).mix <| platformTrace
let args := baseArgs ++ weakArgs ++ traceArgs
let args := baseArgs ++ weakArgs ++ traceArgs ++ lean.ccLinkSharedFlags
let trace ← buildFileUnlessUpToDate dynlib depTrace do
compileSharedLib dynlib args (← getLeanc)
compileSharedLib dynlib args lean.cc
return (dynlib, trace)

/-- Construct a `Dynlib` object for a shared library target. -/
Expand Down
2 changes: 1 addition & 1 deletion src/lake/Lake/Build/Executable.lean
Original file line number Diff line number Diff line change
Expand Up @@ -29,4 +29,4 @@ def LeanExe.recBuildExe (self : LeanExe) : FetchM (BuildJob FilePath) :=
let deps := (← fetch <| self.pkg.facet `deps).push self.pkg
for dep in deps do for lib in dep.externLibs do
linkJobs := linkJobs.push <| ← lib.static.fetch
buildLeanExe self.file linkJobs self.weakLinkArgs self.linkArgs
buildLeanExe self.file linkJobs self.weakLinkArgs self.linkArgs self.sharedLean
4 changes: 4 additions & 0 deletions src/lake/Lake/Build/Trace.lean
Original file line number Diff line number Diff line change
Expand Up @@ -125,6 +125,9 @@ instance : ToString Hash := ⟨Hash.toString⟩
@[inline] def ofByteArray (bytes : ByteArray) : Hash :=
⟨hash bytes⟩

@[inline] def ofBool (b : Bool) :=
mk (hash b)

@[inline] protected def toJson (self : Hash) : Json :=
toJson self.val

Expand All @@ -151,6 +154,7 @@ instance [ComputeHash α m] : ComputeTrace α m Hash := ⟨ComputeHash.computeHa
@[inline] def computeHash [ComputeHash α m] [MonadLiftT m n] (a : α) : n Hash :=
liftM <| ComputeHash.computeHash a

instance : ComputeHash Bool Id := ⟨Hash.ofBool⟩
instance : ComputeHash String Id := ⟨Hash.ofString⟩

/--
Expand Down
49 changes: 34 additions & 15 deletions src/lake/Lake/Config/InstallPath.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,10 +3,12 @@ Copyright (c) 2021 Mac Malone. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mac Malone
-/
import Lean.Compiler.FFI
import Lake.Util.NativeLib
import Lake.Config.Defaults

open System
open System Lean.Compiler.FFI

namespace Lake

/-- Convert the string value of an environment variable to a boolean. -/
Expand Down Expand Up @@ -71,7 +73,13 @@ structure LeanInstall where
initSharedLib := leanSharedLibDir sysroot / initSharedLib
ar : FilePath := "ar"
cc : FilePath := "cc"
customCc : Bool := false
customCc : Bool := true
cFlags := getCFlags sysroot |>.push "-Wno-unused-command-line-argument"
linkStaticFlags := getLinkerFlags sysroot (linkStatic := true)
linkSharedFlags := getLinkerFlags sysroot (linkStatic := false)
ccFlags := cFlags
ccLinkStaticFlags := linkStaticFlags
ccLinkSharedFlags := linkSharedFlags
deriving Inhabited, Repr

/--
Expand All @@ -88,6 +96,10 @@ def LeanInstall.sharedLibPath (self : LeanInstall) : SearchPath :=
def LeanInstall.leanCc? (self : LeanInstall) : Option String :=
if self.customCc then self.cc.toString else none

/-- The link-time flags for the C compiler of the Lean installation. -/
def LeanInstall.ccLinkFlags (shared : Bool) (self : LeanInstall) : Array String :=
if shared then self.ccLinkSharedFlags else self.ccLinkStaticFlags

/-- Lake executable file name. -/
def lakeExe : FilePath :=
FilePath.addExtension "lake" FilePath.exeExtension
Expand Down Expand Up @@ -178,16 +190,15 @@ def LeanInstall.get (sysroot : FilePath) (collocated : Bool := false) : BaseIO L
-- Remark: This is expensive (at least on Windows), so try to avoid it.
getGithash
let ar ← findAr
let (cc, customCc) ← findCc
return {sysroot, githash, ar, cc, customCc}
setCc {sysroot, githash, ar}
where
getGithash := do
EIO.catchExceptions (h := fun _ => pure "") do
let out ← IO.Process.output {
cmd := leanExe sysroot |>.toString,
args := #["--githash"]
}
pure <| out.stdout.trim
return out.stdout.trim
findAr := do
if let some ar ← IO.getEnv "LEAN_AR" then
return FilePath.mk ar
Expand All @@ -199,19 +210,27 @@ where
return ar
else
return "ar"
findCc := do
setCc i := do
if let some cc ← IO.getEnv "LEAN_CC" then
return (FilePath.mk cc, true)
return withCustomCc i cc
else
let cc := leanCcExe sysroot
let cc ←
if (← cc.pathExists) then
pure cc
else if let some cc ← IO.getEnv "CC" then
pure cc
else
pure "cc"
return (cc, false)
if (← cc.pathExists) then
return withInternalCc i cc
else if let some cc ← IO.getEnv "CC" then
return withCustomCc i cc
else
return withCustomCc i "cc"
@[inline] withCustomCc (i : LeanInstall) cc :=
{i with cc}
withInternalCc (i : LeanInstall) cc :=
let ccLinkFlags := getInternalLinkerFlags sysroot
{i with
cc, customCc := false
ccFlags := i.cFlags ++ getInternalCFlags sysroot
ccLinkStaticFlags := ccLinkFlags ++ i.linkStaticFlags
ccLinkSharedFlags := ccLinkFlags ++ i.linkSharedFlags
}

/--
Attempt to detect the installation of the given `lean` command
Expand Down
25 changes: 15 additions & 10 deletions src/lake/Lake/Config/LeanExe.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,7 +3,6 @@ Copyright (c) 2022 Mac Malone. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mac Malone
-/
import Lean.Compiler.FFI
import Lake.Config.Module

namespace Lake
Expand Down Expand Up @@ -74,21 +73,27 @@ The file name of binary executable

/--
The arguments to pass to `leanc` when linking the binary executable.
By default, the package's plus the executable's `moreLinkArgs`.
If `supportInterpreter := true`, Lake links directly to the Lean shared
libraries on Windows by prepending `-leanshared` and adds `-rdynamic` on
other systems.
By default, the package's plus the executable's `moreLinkArgs`.
If `supportInterpreter := true`, Lake prepends `-rdynamic` on non-Windows
systems.
-/
def linkArgs (self : LeanExe) : Array String :=
if self.config.supportInterpreter then
if Platform.isWindows then
#["-leanshared"] ++ self.pkg.moreLinkArgs ++ self.config.moreLinkArgs
else
#["-rdynamic"] ++ self.pkg.moreLinkArgs ++ self.config.moreLinkArgs
if self.config.supportInterpreter && !Platform.isWindows then
#["-rdynamic"] ++ self.pkg.moreLinkArgs ++ self.config.moreLinkArgs
else
self.pkg.moreLinkArgs ++ self.config.moreLinkArgs

/--
Whether the Lean shared library should be dynamically linked to the executable.
If `supportInterpreter := true`, Lean symbols must be visible to the
interpreter. On Windows, it is not possible to statically include these
symbols in the executable due to symbol limits, so Lake dynamically links to
the Lean shared library. Otherwise, Lean is linked statically.
-/
@[inline] def sharedLean (self : LeanExe) : Bool :=
strictAnd Platform.isWindows self.config.supportInterpreter

/--
The arguments to weakly pass to `leanc` when linking the binary executable.
Expand Down
2 changes: 1 addition & 1 deletion src/lake/examples/bootstrap/lakefile.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,5 +8,5 @@ lean_lib Lake

@[default_target]
lean_exe lake where
root := `Lake.Main
root := `LakeMain
supportInterpreter := true

0 comments on commit d4a70f3

Please sign in to comment.