Skip to content

Commit

Permalink
fix: tactics in terms in tactics may break incremental reporting (#4436)
Browse files Browse the repository at this point in the history
A pending tactic mvar managed to escape into an unexpected context in
specific circumstances.

```lean
example : True := by
  · rw [show 0 = 0 by rfl]
```
* Term elaboration of the `show` creates a pending mvar for the `by rfl`
proof
* `rw` fails with an exception because the pattern does not occur in the
target
* `cdot` catches the exception and admits the goal
* `Term.runTactic` [synthesizes all pending mvars from the tactic's
execution](https://github.com/leanprover/lean4/blob/5f9dedfe5ee9972acdebd669f228f487844a6156/src/Lean/Elab/SyntheticMVars.lean#L350),
including the `by rfl` proof. But this would not have happened without
`cdot` as the exception would have skipped that invocation!
* Now incrementality is confused because the nested `by rfl` proof is
unexpectedly run in the same context as the top-level proof, writing to
the wrong promise, and the error message is lost

Solution: disable incrementality for these pending mvars
  • Loading branch information
Kha authored Jun 12, 2024
1 parent bedcbfc commit 8d3be96
Show file tree
Hide file tree
Showing 3 changed files with 25 additions and 1 deletion.
5 changes: 4 additions & 1 deletion src/Lean/Elab/SyntheticMVars.lean
Original file line number Diff line number Diff line change
Expand Up @@ -347,7 +347,10 @@ mutual
-- view even though it is synthetic while a node like `tacticCode` never is (#1990)
withTacticInfoContext tacticCode[0] do
withNarrowedArgTacticReuse (argIdx := 1) (evalTactic ·) tacticCode
synthesizeSyntheticMVars (postpone := .no)
-- Pending tactic mvars may escape from `evalTactic` to here (#4436), so make sure
-- incrementality is disabled so they cannot be confused for top-level tactic blocks
withoutTacticIncrementality true do
synthesizeSyntheticMVars (postpone := .no)
unless remainingGoals.isEmpty do
if report then
reportUnsolvedGoals remainingGoals
Expand Down
9 changes: 9 additions & 0 deletions tests/lean/interactive/incrementalTactic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -71,3 +71,12 @@ def dup_goals : True := by
--^ goals
-- (note that request positions are computed relative to the original document, so the checks above
-- will point at a `show` at run time)

/-!
A tactic mvar may sometimes escape the term elaboration it was created from and should not break
incremental reporting in this case.
-/
-- RESET
def tacInTermInTac : True := by
· rw [show 0 = 0 by rfl]
--^ collectDiagnostics
12 changes: 12 additions & 0 deletions tests/lean/interactive/incrementalTactic.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -62,3 +62,15 @@ s
isRemoved? := none,
hyps := #[] }] }

{"version": 1,
"uri": "file:///incrementalTactic.lean",
"diagnostics":
[{"source": "Lean 4",
"severity": 1,
"range":
{"start": {"line": 2, "character": 8}, "end": {"line": 2, "character": 25}},
"message":
"tactic 'rewrite' failed, did not find instance of the pattern in the target expression\n 0\n⊢ True",
"fullRange":
{"start": {"line": 2, "character": 8},
"end": {"line": 2, "character": 25}}}]}

0 comments on commit 8d3be96

Please sign in to comment.