Skip to content

Commit

Permalink
fix: missing Not propagation rule in grind
Browse files Browse the repository at this point in the history
This PR adds a new propagation rule for negation to the (WIP) `grind` tactic.
  • Loading branch information
leodemoura committed Dec 27, 2024
1 parent c14e5ae commit c8feab5
Show file tree
Hide file tree
Showing 3 changed files with 10 additions and 2 deletions.
4 changes: 4 additions & 0 deletions src/Init/Grind/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -7,6 +7,7 @@ prelude
import Init.Core
import Init.SimpLemmas
import Init.Classical
import Init.ByCases

namespace Lean.Grind

Expand Down Expand Up @@ -41,6 +42,9 @@ theorem not_eq_of_eq_false {a : Prop} (h : a = False) : (Not a) = True := by sim
theorem eq_false_of_not_eq_true {a : Prop} (h : (Not a) = True) : a = False := by simp_all
theorem eq_true_of_not_eq_false {a : Prop} (h : (Not a) = False) : a = True := by simp_all

theorem true_eq_false_of_not_eq_self {a : Prop} (h : (Not a) = a) : True = False := by
by_cases a <;> simp_all

/-! Eq -/

theorem eq_eq_of_eq_true_left {a b : Prop} (h : a = True) : (a = b) = b := by simp [h]
Expand Down
5 changes: 3 additions & 2 deletions src/Lean/Meta/Tactic/Grind/Propagate.lean
Original file line number Diff line number Diff line change
Expand Up @@ -105,12 +105,13 @@ This function performs the following:
- If `(Not a) = True`, propagates `a = False`.
-/
builtin_grind_propagator propagateNotDown ↓Not := fun e => do
let_expr Not a := e | return ()
if (← isEqFalse e) then
let_expr Not a := e | return ()
pushEqTrue a <| mkApp2 (mkConst ``Lean.Grind.eq_true_of_not_eq_false) a (← mkEqFalseProof e)
else if (← isEqTrue e) then
let_expr Not a := e | return ()
pushEqFalse a <| mkApp2 (mkConst ``Lean.Grind.eq_false_of_not_eq_true) a (← mkEqTrueProof e)
else if (← isEqv e a) then
pushEqFalse (← getTrueExpr) <| mkApp2 (mkConst ``Lean.Grind.true_eq_false_of_not_eq_self) a (← mkEqProof e a)

/-- Propagates `Eq` upwards -/
builtin_grind_propagator propagateEqUp ↑Eq := fun e => do
Expand Down
3 changes: 3 additions & 0 deletions tests/lean/run/grind_pre.lean
Original file line number Diff line number Diff line change
Expand Up @@ -124,3 +124,6 @@ example (a : α) (p q r : Prop) : (h₁ : HEq p a) → (h₂ : HEq q a) → (h

example (a b : Nat) (f : Nat → Nat) : (h₁ : a = b) → (h₂ : f a ≠ f b) → False := by
grind

example (a : α) (p q r : Prop) : (h₁ : HEq p a) → (h₂ : HEq q a) → (h₃ : p = r) → q = r := by
grind

0 comments on commit c8feab5

Please sign in to comment.