-
Notifications
You must be signed in to change notification settings - Fork 423
New issue
Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.
By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.
Already on GitHub? Sign in to your account
Running simp?
output produces different goal than simp
#4615
Comments
The issue here is that Here is mwe for (1). inductive Ex (p : α → Prop) : Prop where
| intro (w : α) (h : p w) : Ex p
@[simp] theorem ex_prop (p q : Prop) : Ex (fun h : p => q) = (p ∧ q) := sorry
@[simp high] theorem ex_and_left (p : α → Prop) (b : Prop) : (Ex fun x => b ∧ p x) = (b ∧ Ex fun x => p x) := sorry
@[simp] theorem ex_eq_right_right (p q : α → Prop) : (Ex fun x => p x ∧ q x ∧ x = a) = (p a ∧ q a) := sorry
/--
info: Try this: simp only [ex_prop, ex_and_left, ex_eq_right_right]
---
warning: declaration uses 'sorry'
-/
#guard_msgs in
example (x : Nat) : Ex fun (a : Nat) => Ex fun (_ : 0 < a) => Ex fun (_ : 1 < a) => a = x := by
simp? -- Produces incorrect `simp only`
guard_target =ₛ 1 < x ∧ 0 < x
sorry
example (x : Nat) : Ex fun (a : Nat) => Ex fun (_ : 0 < a) => Ex fun (_ : 1 < a) => a = x := by
simp only [ex_prop, ex_and_left, ex_eq_right_right]
guard_target =ₛ 0 < x ∧ 1 < x
sorry
example (x : Nat) : Ex fun (a : Nat) => Ex fun (_ : 0 < a) => Ex fun (_ : 1 < a) => a = x := by
simp only [ex_and_left, ex_prop, ex_eq_right_right] -- This is the correct order
guard_target =ₛ 1 < x ∧ 0 < x
sorry The original issue is a mwe for (2). Here is another mwe for (2) that does not depend on the order theorems are defined in inductive Ex (p : α → Prop) : Prop where
| intro (w : α) (h : p w) : Ex p
@[simp] theorem ex_and_left (p : α → Prop) (b : Prop) : (Ex fun x => b ∧ p x) = (b ∧ Ex fun x => p x) := sorry
@[simp] theorem ex_prop (p q : Prop) : Ex (fun h : p => q) = (p ∧ q) := sorry
@[simp] theorem ex_eq_right_right (p q : α → Prop) : (Ex fun x => p x ∧ q x ∧ x = a) = (p a ∧ q a) := sorry
/--
info: Try this: simp only [ex_prop, ex_and_left, ex_eq_right_right]
---
warning: declaration uses 'sorry'
-/
#guard_msgs in
example (x : Nat) : Ex fun (a : Nat) => Ex fun (_ : 0 < a) => Ex fun (_ : 1 < a) => a = x := by
simp?
guard_target =ₛ 1 < x ∧ 0 < x
sorry
example (x : Nat) : Ex fun (a : Nat) => Ex fun (_ : 0 < a) => Ex fun (_ : 1 < a) => a = x := by
simp only [ex_prop, ex_and_left, ex_eq_right_right]
guard_target =ₛ 0 < x ∧ 1 < x
sorry
example (x : Nat) : Ex fun (a : Nat) => Ex fun (_ : 0 < a) => Ex fun (_ : 1 < a) => a = x := by
simp only [ex_and_left, ex_prop, ex_eq_right_right] -- This is the correct order
guard_target =ₛ 1 < x ∧ 0 < x
sorry |
inductive Ex (p : α → Prop) : Prop where
| intro (w : α) (h : p w) : Ex p
@[seval] theorem ex_prop (p q : Prop) : Ex (fun h : p => q) = (p ∧ q) := sorry
@[simp] theorem ex_and_left (p : α → Prop) (b : Prop) : (Ex fun x => b ∧ p x) = (b ∧ Ex fun x => p x) := sorry
@[simp] theorem ex_eq_right_right (p q : α → Prop) : (Ex fun x => p x ∧ q x ∧ x = a) = (p a ∧ q a) := sorry
/--
info: Try this: simp only [ex_prop, ex_and_left, ex_eq_right_right]
---
warning: declaration uses 'sorry'
-/
#guard_msgs in
example (x : Nat) : Ex fun (a : Nat) => Ex fun (_ : 0 < a) => Ex fun (_ : 1 < a) => a = x := by
-- If theorems in `simp` and `seval` have the same priority, the one in `simp` must appear first
-- in the resulting `simp only`
simp? only [simp, seval, -Nat.isValue] -- produces an incorrect `simp only`
guard_target =ₛ 1 < x ∧ 0 < x
sorry
example (x : Nat) : Ex fun (a : Nat) => Ex fun (_ : 0 < a) => Ex fun (_ : 1 < a) => a = x := by
simp only [ex_prop, ex_and_left, ex_eq_right_right]
guard_target =ₛ 0 < x ∧ 1 < x
sorry
example (x : Nat) : Ex fun (a : Nat) => Ex fun (_ : 0 < a) => Ex fun (_ : 1 < a) => a = x := by
simp only [ex_and_left, ex_eq_right_right, ex_prop] -- This is the correct order
guard_target =ₛ 1 < x ∧ 0 < x
sorry |
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
Context
Originally reported on Zulip.
Steps to Reproduce
Paste the above code into a Lean editor.
Expected behavior: The goal is the same.
Actual behavior: The goal is different.
Versions
This particular repro works on 4.7.0+ (where the simp lemmas were introduced), but there may be a simpler reproduction that uses pre-4.7.0 simp lemmas.
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: