You signed in with another tab or window. Reload to refresh your session.You signed out in another tab or window. Reload to refresh your session.You switched accounts on another tab or window. Reload to refresh your session.Dismiss alert
Lean.MVarId.revertAll fails to ignore auxiliary declarations if variable is used.
Context
[Broader context that the issue occurred in. If there was any prior discussion on the Lean Zulip, link it here as well.]
Steps to Reproduce
import Lean
open Lean.Elab.Tactic
variable (p q : Prop)
theoremfoo (h : p ∧ q) : q ∧ p := by
run_tac liftMetaTactic1 (·.revertAll)
sorry
Expected behavior: Goal is ⊢ ∀ (p q : Prop), p ∧ q → q ∧ p, no errors
Actual behavior: Goal is ⊢ ∀ (p q : Prop), p ∧ q → (p ∧ q → q ∧ p) → q ∧ p, and the error is
fail to show termination for
foo
with errors
failed to infer structural recursion:
Not considering parameter p of foo:
it is unchanged in the recursive calls
Not considering parameter q of foo:
it is unchanged in the recursive calls
Cannot use parameter h:
the type p ∧ q does not have a `.brecOn` recursor
Versions
4.14.0-rc3
Additional Information
[Additional information, configuration or data that might be necessary to reproduce the issue]
Empirically, the current behavior seems to be "ignore leading auxiliary declarations", so I'd lean towards this being a design issue rather than a documentation issue.
Prerequisites
Please put an X between the brackets as you perform the following steps:
https://github.com/leanprover/lean4/issues
Avoid dependencies to Mathlib or Batteries.
https://live.lean-lang.org/#project=lean-nightly
(You can also use the settings there to switch to “Lean nightly”)
Description
Lean.MVarId.revertAll
fails to ignore auxiliary declarations ifvariable
is used.Context
[Broader context that the issue occurred in. If there was any prior discussion on the Lean Zulip, link it here as well.]
Steps to Reproduce
Expected behavior: Goal is
⊢ ∀ (p q : Prop), p ∧ q → q ∧ p
, no errorsActual behavior: Goal is
⊢ ∀ (p q : Prop), p ∧ q → (p ∧ q → q ∧ p) → q ∧ p
, and the error isVersions
4.14.0-rc3
Additional Information
[Additional information, configuration or data that might be necessary to reproduce the issue]
Impact
Add 👍 to issues you consider important. If others are impacted by this issue, please ask them to add 👍 to it.
The text was updated successfully, but these errors were encountered: