Skip to content

Commit

Permalink
feat: asynchronous code generation
Browse files Browse the repository at this point in the history
  • Loading branch information
Kha committed Feb 3, 2025
1 parent d68c2ce commit c47e9a9
Show file tree
Hide file tree
Showing 11 changed files with 59 additions and 77 deletions.
5 changes: 5 additions & 0 deletions src/Lean/Compiler/IR/CompilerM.lean
Original file line number Diff line number Diff line change
Expand Up @@ -89,6 +89,11 @@ builtin_initialize declMapExt : SimplePersistentEnvExtension Decl DeclMap ←
toArrayFn := fun s =>
let decls := s.foldl (init := #[]) fun decls decl => decls.push decl
sortDecls decls
-- Written to on codegen environment branch but accessed from other elaboration branches when
-- calling into the interpreter. We cannot use `async` as the IR declarations added may not
-- share a name prefix with the top-level Lean declaration being compiled, e.g. from
-- specialization.
asyncMode := .sync
}

@[export lean_ir_find_env_decl]
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Compiler/LCNF/AuxDeclCache.lean
Original file line number Diff line number Diff line change
Expand Up @@ -12,7 +12,8 @@ namespace Lean.Compiler.LCNF

abbrev AuxDeclCache := PHashMap Decl Name

builtin_initialize auxDeclCacheExt : EnvExtension AuxDeclCache ← registerEnvExtension (pure {})
builtin_initialize auxDeclCacheExt : EnvExtension AuxDeclCache ←
registerEnvExtension (pure {}) (asyncMode := .sync) -- compilation is non-parallel anyway

inductive CacheAuxDeclResult where
| new
Expand All @@ -29,4 +30,3 @@ def cacheAuxDecl (decl : Decl) : CompilerM CacheAuxDeclResult := do
return .new

end Lean.Compiler.LCNF

1 change: 1 addition & 0 deletions src/Lean/Compiler/LCNF/PhaseExt.lean
Original file line number Diff line number Diff line change
Expand Up @@ -32,6 +32,7 @@ def mkDeclExt (name : Name := by exact decl_name%) : IO DeclExt := do
exportEntriesFn := fun s =>
let decls := s.foldl (init := #[]) fun decls _ decl => decls.push decl
sortDecls decls
asyncMode := .sync -- compilation is non-parallel anyway
}

builtin_initialize baseExt : PersistentEnvExtension Decl Decl DeclExtState ← mkDeclExt
Expand Down
1 change: 1 addition & 0 deletions src/Lean/Compiler/Specialize.lean
Original file line number Diff line number Diff line change
Expand Up @@ -110,6 +110,7 @@ builtin_initialize specExtension : SimplePersistentEnvExtension SpecEntry SpecSt
registerSimplePersistentEnvExtension {
addEntryFn := SpecState.addEntry,
addImportedFn := fun es => (mkStateFromImportedEntries SpecState.addEntry {} es).switch
asyncMode := .sync -- compilation is non-parallel anyway
}

@[export lean_add_specialization_info]
Expand Down
50 changes: 29 additions & 21 deletions src/Lean/CoreM.lean
Original file line number Diff line number Diff line change
Expand Up @@ -521,39 +521,47 @@ opaque compileDeclsNew (declNames : List Name) : CoreM Unit
@[extern "lean_compile_decls"]
opaque compileDeclsOld (env : Environment) (opt : @& Options) (decls : @& List Name) : Except Kernel.Exception Environment

def compileDecl (decl : Declaration) : CoreM Unit := do
-- `ref?` is used for error reporting if available
partial def compileDecls (decls : List Name) (ref? : Option Declaration := none)
(logErrors := true) : CoreM Unit := do
if !Elab.async.get (← getOptions) then
doCompile
return
let env ← getEnv
let (postEnv, prom) ← env.promiseChecked
let checkAct ← Core.wrapAsyncAsSnapshot fun _ => do
try
doCompile
finally
prom.resolve (← getEnv)
let t ← BaseIO.mapTask (fun _ => checkAct) env.checked
let endRange? := (← getRef).getTailPos?.map fun pos => ⟨pos, pos⟩
Core.logSnapshotTask { range? := endRange?, task := t }
setEnv postEnv
where doCompile := do
-- don't compile if kernel errored; should be converted into a task dependency when compilation
-- is made async as well
if !decl.getNames.all (← getEnv).constants.contains then
if !decls.all (← getEnv).constants.contains then
return
let opts ← getOptions
let decls := Compiler.getDeclNamesForCodeGen decl
if compiler.enableNew.get opts then
compileDeclsNew decls

let res ← withTraceNode `compiler (fun _ => return m!"compiling old: {decls}") do
return compileDeclsOld (← getEnv) opts decls
match res with
| Except.ok env => setEnv env
| Except.error (.other msg) =>
checkUnsupported decl -- Generate nicer error message for unsupported recursors and axioms
throwError msg
| Except.error ex =>
throwKernelException ex

def compileDecls (decls : List Name) : CoreM Unit := do
-- don't compile if kernel errored; should be converted into a task dependency when compilation
-- is made async as well
if !decls.all (← getEnv).constants.contains then
return
let opts ← getOptions
if compiler.enableNew.get opts then
compileDeclsNew decls
match compileDeclsOld (← getEnv) opts decls with
| Except.ok env => setEnv env
| Except.error (.other msg) =>
throwError msg
if logErrors then
if let some decl := ref? then
checkUnsupported decl -- Generate nicer error message for unsupported recursors and axioms
throwError msg
| Except.error ex =>
throwKernelException ex
if logErrors then
throwKernelException ex

def compileDecl (decl : Declaration) (logErrors := true) : CoreM Unit := do
compileDecls (Compiler.getDeclNamesForCodeGen decl) decl logErrors

def getDiag (opts : Options) : Bool :=
diagnostics.get opts
Expand Down
15 changes: 4 additions & 11 deletions src/Lean/Elab/PreDefinition/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -102,15 +102,8 @@ def addAsAxiom (preDef : PreDefinition) : MetaM Unit := do
private def shouldGenCodeFor (preDef : PreDefinition) : Bool :=
!preDef.kind.isTheorem && !preDef.modifiers.isNoncomputable

private def compileDecl (decl : Declaration) : TermElabM Bool := do
try
Lean.compileDecl decl
catch ex =>
if (← read).isNoncomputableSection then
return false
else
throw ex
return true
private def compileDecl (decl : Declaration) : TermElabM Unit := do
Lean.compileDecl (logErrors := !(← read).isNoncomputableSection) decl

register_builtin_option diagnostics.threshold.proofSize : Nat := {
defValue := 16384
Expand Down Expand Up @@ -166,7 +159,7 @@ private def addNonRecAux (preDef : PreDefinition) (compile : Bool) (all : List N
if preDef.modifiers.isNoncomputable then
modifyEnv fun env => addNoncomputable env preDef.declName
if compile && shouldGenCodeFor preDef then
discard <| compileDecl decl
compileDecl decl
if applyAttrAfterCompilation then
generateEagerEqns preDef.declName
applyAttributesOf #[preDef] AttributeApplicationTime.afterCompilation
Expand Down Expand Up @@ -206,7 +199,7 @@ def addAndCompileUnsafe (preDefs : Array PreDefinition) (safety := DefinitionSaf
for preDef in preDefs do
addTermInfo' preDef.ref (← mkConstWithLevelParams preDef.declName) (isBinder := true)
applyAttributesOf preDefs AttributeApplicationTime.afterTypeChecking
discard <| compileDecl decl
compileDecl decl
applyAttributesOf preDefs AttributeApplicationTime.afterCompilation
return ()

Expand Down
13 changes: 7 additions & 6 deletions src/Lean/Elab/PreDefinition/Main.lean
Original file line number Diff line number Diff line change
Expand Up @@ -300,12 +300,13 @@ def addPreDefinitions (preDefs : Array PreDefinition) : TermElabM Unit := withLC
try
if preDefs.all fun preDef => (preDef.kind matches DefKind.def | DefKind.instance) || preDefs.all fun preDef => preDef.kind == DefKind.abbrev then
-- try to add as partial definition
try
addAndCompilePartial preDefs (useSorry := true)
catch _ =>
-- Compilation failed try again just as axiom
s.restore
addAsAxioms preDefs
withOptions (Elab.async.set · false) do
try
addAndCompilePartial preDefs (useSorry := true)
catch _ =>
-- Compilation failed try again just as axiom
s.restore
addAsAxioms preDefs
else if preDefs.all fun preDef => preDef.kind == DefKind.theorem then
addAsAxioms preDefs
catch _ => s.restore
Expand Down
11 changes: 7 additions & 4 deletions src/Lean/Environment.lean
Original file line number Diff line number Diff line change
Expand Up @@ -451,9 +451,7 @@ def ofKernelEnv (env : Kernel.Environment) : Environment :=

@[export lean_elab_environment_to_kernel_env]
def toKernelEnv (env : Environment) : Kernel.Environment :=
-- TODO: should just be the following when we store extension data in `checked`
--env.checked.get
{ env.checked.get with extensions := env.checkedWithoutAsync.extensions }
env.checked.get

/-- Consistently updates synchronous and asynchronous parts of the environment without blocking. -/
private def modifyCheckedAsync (env : Environment) (f : Kernel.Environment → Kernel.Environment) : Environment :=
Expand All @@ -463,6 +461,10 @@ private def modifyCheckedAsync (env : Environment) (f : Kernel.Environment → K
private def setCheckedSync (env : Environment) (newChecked : Kernel.Environment) : Environment :=
{ env with checked := .pure newChecked, checkedWithoutAsync := newChecked }

def promiseChecked (env : Environment) : BaseIO (Environment × IO.Promise Environment) := do
let prom ← IO.Promise.new
return ({ env with checked := prom.result.bind (sync := true) (·.checked) }, prom)

/--
Checks whether the given declaration name may potentially added, or have been added, to the current
environment branch, which is the case either if this is the main branch or if the declaration name
Expand Down Expand Up @@ -621,6 +623,7 @@ information.
-/
def addConstAsync (env : Environment) (constName : Name) (kind : ConstantKind) (reportExts := true) :
IO AddConstAsyncResult := do
assert! env.asyncMayContain constName
let sigPromise ← IO.Promise.new
let infoPromise ← IO.Promise.new
let extensionsPromise ← IO.Promise.new
Expand Down Expand Up @@ -1565,7 +1568,7 @@ def getNamespaceSet (env : Environment) : NameSSet :=

@[export lean_elab_environment_update_base_after_kernel_add]
private def updateBaseAfterKernelAdd (env : Environment) (kernel : Kernel.Environment) : Environment :=
env.setCheckedSync { kernel with extensions := env.checkedWithoutAsync.extensions }
{ env with checked := .pure kernel, checkedWithoutAsync := { kernel with extensions := env.checkedWithoutAsync.extensions } }

@[export lean_display_stats]
def displayStats (env : Environment) : IO Unit := do
Expand Down
1 change: 1 addition & 0 deletions tests/lean/computedFieldsCode.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -54,6 +54,7 @@ def Exp.casesOn._override._rarg._boxed (x_1 : obj) (x_2 : obj) (x_3 : obj) (x_4
dec x_5;
dec x_4;
ret x_9

[result]
def Exp.var._override (x_1 : u32) : obj :=
let x_2 : u64 := UInt32.toUInt64 x_1;
Expand Down
31 changes: 0 additions & 31 deletions tests/lean/ll_infer_type_bug.lean

This file was deleted.

4 changes: 2 additions & 2 deletions tests/lean/run/2291.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,6 +8,8 @@ The following example would cause the pretty printer to panic.

set_option trace.compiler.simp true in
/--
info: [0]
---
info: [compiler.simp] >> _eval
let _x_21 := `Nat;
let _x_22 := [];
Expand Down Expand Up @@ -105,8 +107,6 @@ let _x_1 :=
Lean.List.toExprAux._at._eval._spec_1✝ _eval._closed_9 _eval._closed_13
_eval._closed_14;
Lean.MessageData.ofExpr _x_1
---
info: [0]
-/
#guard_msgs in
#eval [0]
Expand Down

0 comments on commit c47e9a9

Please sign in to comment.