Skip to content

Commit

Permalink
chore: clean up Elab.async handling (#6299)
Browse files Browse the repository at this point in the history
* Make sure metaprogramming users cannot be surprised by its
introduction
* Make `#guard_msgs` compatible with its use
  • Loading branch information
Kha authored Dec 3, 2024
1 parent 473274f commit 00718c3
Show file tree
Hide file tree
Showing 4 changed files with 27 additions and 10 deletions.
12 changes: 10 additions & 2 deletions src/Lean/CoreM.lean
Original file line number Diff line number Diff line change
Expand Up @@ -32,8 +32,16 @@ register_builtin_option maxHeartbeats : Nat := {
}

register_builtin_option Elab.async : Bool := {
defValue := true
descr := "perform elaboration using multiple threads where possible"
defValue := false
descr := "perform elaboration using multiple threads where possible\
\n\
\nThis option defaults to `false` but (when not explicitly set) is overridden to `true` in \
`Lean.Language.Lean.process` as used by the cmdline driver and language server. \
Metaprogramming users driving elaboration directly via e.g. \
`Lean.Elab.Command.elabCommandTopLevel` can opt into asynchronous elaboration by setting \
this option but then are responsible for processing messages and other data not only in the \
resulting command state but also from async tasks in `Lean.Command.Context.snap?` and \
`Lean.Command.State.snapshotTasks`."
}

/--
Expand Down
7 changes: 6 additions & 1 deletion src/Lean/Elab/Command.lean
Original file line number Diff line number Diff line change
Expand Up @@ -101,7 +101,7 @@ structure Context where
(mutual) defs and contained tactics, in which case the `DynamicSnapshot` is a
`HeadersParsedSnapshot`.
Definitely resolved in `Language.Lean.process.doElab`.
Definitely resolved in `Lean.Elab.Command.elabCommandTopLevel`.
Invariant: if the bundle's `old?` is set, the context and state at the beginning of current and
old elaboration are identical.
Expand Down Expand Up @@ -562,6 +562,11 @@ def elabCommandTopLevel (stx : Syntax) : CommandElabM Unit := withRef stx do pro
withLogging do
runLintersAsync stx
finally
-- Make sure `snap?` is definitely resolved; we do not use it for reporting as `#guard_msgs` may
-- be the caller of this function and add new messages and info trees
if let some snap := (← read).snap? then
snap.new.resolve default

-- note the order: first process current messages & info trees, then add back old messages & trees,
-- then convert new traces to messages
let mut msgs := (← get).messages
Expand Down
16 changes: 9 additions & 7 deletions src/Lean/Elab/GuardMsgs.lean
Original file line number Diff line number Diff line change
Expand Up @@ -140,13 +140,15 @@ def MessageOrdering.apply (mode : MessageOrdering) (msgs : List String) : List S
|>.trim |> removeTrailingWhitespaceMarker
let (whitespace, ordering, specFn) ← parseGuardMsgsSpec spec?
let initMsgs ← modifyGet fun st => (st.messages, { st with messages := {} })
-- disable async elaboration for `cmd` so `msgs` will contain all messages
let async := Elab.async.get (← getOptions)
modifyScope fun scope => { scope with opts := Elab.async.set scope.opts false }
-- The `#guard_msgs` command is special-cased in `elabCommandTopLevel` to ensure linters only run once.
elabCommandTopLevel cmd
modifyScope fun scope => { scope with opts := Elab.async.set scope.opts async }
let msgs := (← get).messages
-- do not forward snapshot as we don't want messages assigned to it to leak outside
withReader ({ · with snap? := none }) do
-- The `#guard_msgs` command is special-cased in `elabCommandTopLevel` to ensure linters only run once.
elabCommandTopLevel cmd
-- collect sync and async messages
let msgs := (← get).messages ++
(← get).snapshotTasks.foldl (· ++ ·.get.getAll.foldl (· ++ ·.diagnostics.msgLog) {}) {}
-- clear async messages as we don't want them to leak outside
modify ({ · with snapshotTasks := #[] })
let mut toCheck : MessageLog := .empty
let mut toPassthrough : MessageLog := .empty
for msg in msgs.toList do
Expand Down
2 changes: 2 additions & 0 deletions src/Lean/Language/Lean.lean
Original file line number Diff line number Diff line change
Expand Up @@ -433,6 +433,8 @@ where
}
-- now that imports have been loaded, check options again
let opts ← reparseOptions setup.opts
-- default to async elaboration; see also `Elab.async` docs
let opts := Elab.async.setIfNotSet opts true
let cmdState := Elab.Command.mkState headerEnv msgLog opts
let cmdState := { cmdState with
infoState := {
Expand Down

0 comments on commit 00718c3

Please sign in to comment.