Skip to content

Commit

Permalink
fix: lake: track trace of cached build logs (#4343)
Browse files Browse the repository at this point in the history
Stores the dependency trace for a build in the cached build log and then
verifies that it matches the trace of the current build before replaying
the log. Includes test.

Closes #4303.
  • Loading branch information
tydeu committed Jun 5, 2024
1 parent 1da9377 commit df668f0
Show file tree
Hide file tree
Showing 9 changed files with 69 additions and 13 deletions.
26 changes: 19 additions & 7 deletions src/lake/Lake/Build/Common.lean
Original file line number Diff line number Diff line change
Expand Up @@ -95,27 +95,39 @@ def cacheFileHash (file : FilePath) : IO Hash := do
IO.FS.writeFile hashFile hash.toString
return hash

/-- The build log file format. -/
structure BuildLog where
depHash : Hash
log : Log
deriving ToJson, FromJson

/--
Replays the JSON log in `logFile` (if it exists).
If the log file is malformed, logs a warning.
-/
def replayBuildLog (logFile : FilePath) : LogIO PUnit := do
def replayBuildLog (logFile : FilePath) (depTrace : BuildTrace) : LogIO PUnit := do
match (← IO.FS.readFile logFile |>.toBaseIO) with
| .ok contents =>
match Json.parse contents >>= fromJson? with
| .ok entries => Log.mk entries |>.replay
| .error e => logWarning s!"cached build log is corrupted: {e}"
| .ok {log, depHash : BuildLog} =>
if depTrace.hash == depHash then
log.replay
| .error e => logWarning s!"failed to read cached build log: {e}"
| .error (.noFileOrDirectory ..) => pure ()
| .error e => logWarning s!"failed to read cached build log: {e}"

/-- Saves the log produce by `build` as JSON to `logFile`. -/
def cacheBuildLog (logFile : FilePath) (build : JobM PUnit) : JobM PUnit := do
def cacheBuildLog
(logFile : FilePath) (depTrace : BuildTrace) (build : JobM PUnit)
: JobM PUnit := do
let iniPos ← getLogPos
let errPos? ← try build; pure none catch errPos => pure (some errPos)
let log := (← getLog).takeFrom iniPos
unless log.isEmpty do
let log := {log, depHash := depTrace.hash : BuildLog}
IO.FS.writeFile logFile (toJson log).pretty
if let some errPos := errPos? then throw errPos
if let some errPos := errPos? then
throw errPos

/--
Builds `file` using `build` unless it already exists and `depTrace` matches
Expand All @@ -133,10 +145,10 @@ def buildFileUnlessUpToDate
: JobM BuildTrace := do
let traceFile := FilePath.mk <| file.toString ++ ".trace"
let logFile := FilePath.mk <| file.toString ++ ".log.json"
let build := cacheBuildLog logFile build
let build := cacheBuildLog logFile depTrace build
if (← buildUnlessUpToDate? file depTrace traceFile build) then
updateAction .replay
replayBuildLog logFile
replayBuildLog logFile depTrace
fetchFileTrace file
else
return .mk (← cacheFileHash file) (← getMTime file)
Expand Down
4 changes: 2 additions & 2 deletions src/lake/Lake/Build/Module.lean
Original file line number Diff line number Diff line change
Expand Up @@ -157,7 +157,7 @@ def Module.recBuildLean (mod : Module) : FetchM (BuildJob Unit) := do
let upToDate ← buildUnlessUpToDate? (oldTrace := srcTrace) mod modTrace mod.traceFile do
let hasLLVM := Lean.Internal.hasLLVMBackend ()
let bcFile? := if hasLLVM then some mod.bcFile else none
cacheBuildLog mod.logFile do
cacheBuildLog mod.logFile modTrace do
compileLeanModule mod.leanFile mod.oleanFile mod.ileanFile mod.cFile bcFile?
(← getLeanPath) mod.rootDir dynlibs dynlibPath (mod.weakLeanArgs ++ mod.leanArgs) (← getLean)
discard <| cacheFileHash mod.oleanFile
Expand All @@ -167,7 +167,7 @@ def Module.recBuildLean (mod : Module) : FetchM (BuildJob Unit) := do
discard <| cacheFileHash mod.bcFile
if upToDate then
updateAction .replay
replayBuildLog mod.logFile
replayBuildLog mod.logFile modTrace
return ((), depTrace)

/-- The `ModuleFacetConfig` for the builtin `leanArtsFacet`. -/
Expand Down
13 changes: 12 additions & 1 deletion src/lake/Lake/Build/Trace.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,8 +5,9 @@ Authors: Mac Malone
-/
import Lake.Util.IO
import Lake.Util.Newline
import Lean.Data.Json

open System
open System Lean

namespace Lake

Expand Down Expand Up @@ -111,6 +112,16 @@ instance : ToString Hash := ⟨Hash.toString⟩
@[inline] def ofByteArray (bytes : ByteArray) : Hash :=
⟨hash bytes⟩

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

instance : ToJson Hash := ⟨Hash.toJson⟩

@[inline] protected def fromJson? (json : Json) : Except String Hash :=
(⟨·⟩) <$> fromJson? json

instance : FromJson Hash := ⟨Hash.fromJson?⟩

end Hash

class ComputeHash (α : Type u) (m : outParam $ TypeType v) where
Expand Down
3 changes: 0 additions & 3 deletions src/lake/Lake/Load/Elab.lean
Original file line number Diff line number Diff line change
Expand Up @@ -142,9 +142,6 @@ where
-- IR Extension (for constant evaluation)
|>.insert ``IR.declMapExt

instance : ToJson Hash := ⟨(toJson ·.val)⟩
instance : FromJson Hash := ⟨((⟨·⟩) <$> fromJson? ·)⟩

structure ConfigTrace where
platform : String
leanHash : String
Expand Down
1 change: 1 addition & 0 deletions src/lake/tests/reversion/Hello.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
def hello := "foo"
4 changes: 4 additions & 0 deletions src/lake/tests/reversion/Main.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,4 @@
import Hello

def main (args : List String) : IO Unit :=
IO.println s!"Hello, {hello}!"
1 change: 1 addition & 0 deletions src/lake/tests/reversion/clean.sh
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
rm -rf .lake lake-manifest.json Hello.lean
10 changes: 10 additions & 0 deletions src/lake/tests/reversion/lakefile.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,10 @@
import Lake
open Lake DSL

package test

lean_lib Hello

@[default_target]
lean_exe hello where
root := `Main
20 changes: 20 additions & 0 deletions src/lake/tests/reversion/test.sh
Original file line number Diff line number Diff line change
@@ -0,0 +1,20 @@
#!/usr/bin/env bash
set -euxo pipefail

LAKE=${LAKE:-../../.lake/build/bin/lake}

./clean.sh

# Test that introducing an error and reverting works
# https://github.com/leanprover/lean4/issues/4303

# Initial state
echo 'def hello := "foo"' > Hello.lean
$LAKE -q build
# Introduce error
echo 'error' > Hello.lean
$LAKE build && exit 1 || true
# Revert
echo 'def hello := "foo"' > Hello.lean
# Ensure error is not presevered but the warning in another file is
$LAKE -q build | grep --color 'Replayed Main'

0 comments on commit df668f0

Please sign in to comment.