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
Reduce the issue to a minimal, self-contained, reproducible test case. Avoid dependencies to mathlib4 or std4.
Note this issue may be related to #2281, I am not sure.
Description
simp (config :={zeta := false}) sometimes unfold let values.
Context
This is a regression compared to Lean 3 that complicates porting the sphere eversion project for instance. It was discussed on Zulip.
Steps to Reproduce
deff : Nat → Nat := fun x ↦ x - x
@[simp]theoremf_zero (n : Nat) : f n = 0 := Nat.sub_self n
example (n : Nat) : False := bylet g := f n
have : g + n = n := by
simp (config := {zeta := false}) [Nat.zero_add]
sorry
Expected behavior:
The simp should fail. In Lean 3 the dsimp line below is needed.
-- Lean 3 codedeff : nat → nat := fun x, x - x
@[simp]lemmaf_zero (n : ℕ) : f n = 0 := nat.sub_self n
example (n : nat) : false :=
beginlet g := f n,
have : g + n = n,
{ -- dsimp only [g],
simp [nat.zero_add] },
sorryend
Actual behavior:
Lean 4 happily closes the goal. Of course in my example this looks like a feature, but in actual examples it prevents simp from closing the goal, by sending it to a dead end.
This commit also removes parameter `simpleReduce` from discrimination
trees, and take WHNF configuration options.
Reason: it is more dynamic now. For example, the simplifier
will be able to use different configurations for discrimination tree insertion
and retrieval. We need this feature to address issues #2669 and #2281
This commit also removes the dead Meta.Config field `zetaNonDep`.
This commit also removes parameter `simpleReduce` from discrimination
trees, and take WHNF configuration options.
Reason: it is more dynamic now. For example, the simplifier
will be able to use different configurations for discrimination tree insertion
and retrieval. We need this feature to address issues #2669 and #2281
This commit also removes the dead Meta.Config field `zetaNonDep`.
Prerequisites
Note this issue may be related to #2281, I am not sure.
Description
simp (config :={zeta := false})
sometimes unfold let values.Context
This is a regression compared to Lean 3 that complicates porting the sphere eversion project for instance. It was discussed on Zulip.
Steps to Reproduce
Expected behavior:
The simp should fail. In Lean 3 the
dsimp
line below is needed.Actual behavior:
Lean 4 happily closes the goal. Of course in my example this looks like a feature, but in actual examples it prevents
simp
from closing the goal, by sending it to a dead end.Versions
Lean (version 4.2.0-rc1, commit a62d2fd, Release)
Ubuntu 22.04
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: