Skip to content
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

Open
3 tasks done
llllvvuu opened this issue Jul 2, 2024 · 2 comments
Open
3 tasks done

Running simp? output produces different goal than simp #4615

llllvvuu opened this issue Jul 2, 2024 · 2 comments
Labels
bug Something isn't working P-medium We may work on this issue if we find the time

Comments

@llllvvuu
Copy link
Contributor

llllvvuu commented Jul 2, 2024

Prerequisites

Please put an X between the brackets as you perform the following steps:

Description

 /-- info: "4.10.0-rc1" -/
#guard_msgs in
#eval Lean.versionString

/--
info: Try this: simp only [exists_prop, exists_and_left, exists_eq_right_right]
---
warning: declaration uses 'sorry'
---
info: [Meta.Tactic.simp.rewrite] exists_prop:1000, ∃ x_1, a = x ==> 1 < a ∧ a = x
[Meta.Tactic.simp.rewrite] exists_and_left:1000, ∃ x_1, 1 < a ∧ a = x ==> 1 < a ∧ ∃ x_1, a = x
[Meta.Tactic.simp.rewrite] exists_prop:1000, ∃ x_1, a = x ==> 0 < a ∧ a = x
[Meta.Tactic.simp.rewrite] exists_eq_right_right:1000, ∃ a, 1 < a ∧ 0 < a ∧ a = x ==> 1 < x ∧ 0 < x
-/
#guard_msgs in
set_option trace.Meta.Tactic.simp.rewrite true in
example (x : Nat) : ∃ (a : Nat) (_ : 0 < a) (_ : 1 < a), a = x := by
  simp?
  show 1 < x ∧ 0 < x
  sorry

example (x : Nat) : ∃ (a : Nat) (_ : 0 < a) (_ : 1 < a), a = x := by
  simp only [exists_prop, exists_and_left, exists_eq_right_right]
  show 0 < x ∧ 1 < x
  sorry

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.

@llllvvuu llllvvuu added the bug Something isn't working label Jul 2, 2024
@leodemoura
Copy link
Member

The issue here is that simp? produces a simp only [...] using the order the theorems have been applied.
The correct order should be based on
1- [simp] attribute priorities
2- The order the theorems have been inserted into the global simp-set.

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 Init

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

@leodemoura
Copy link
Member

leodemoura commented Jul 2, 2024

simp can take additional simp-sets, the order we provide them should be taken into account too. Here is a mwe for this related issue.

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

@leanprover-bot leanprover-bot added the P-medium We may work on this issue if we find the time label Aug 2, 2024
Sign up for free to join this conversation on GitHub. Already have an account? Sign in to comment
Labels
bug Something isn't working P-medium We may work on this issue if we find the time
Projects
None yet
Development

No branches or pull requests

3 participants