Skip to content

Commit

Permalink
feat: add support for splitting on <-> to grind
Browse files Browse the repository at this point in the history
This PR adds support for case-splitting on `<->` (and `@Eq Prop`) in
the `grind` tactic.
  • Loading branch information
leodemoura committed Jan 12, 2025
1 parent acad587 commit f94fc31
Show file tree
Hide file tree
Showing 7 changed files with 80 additions and 12 deletions.
6 changes: 6 additions & 0 deletions src/Init/Grind/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -66,6 +66,12 @@ theorem eq_eq_of_eq_true_right {a b : Prop} (h : b = True) : (a = b) = a := by s
theorem eq_congr {α : Sort u} {a₁ b₁ a₂ b₂ : α} (h₁ : a₁ = a₂) (h₂ : b₁ = b₂) : (a₁ = b₁) = (a₂ = b₂) := by simp [*]
theorem eq_congr' {α : Sort u} {a₁ b₁ a₂ b₂ : α} (h₁ : a₁ = b₂) (h₂ : b₁ = a₂) : (a₁ = b₁) = (a₂ = b₂) := by rw [h₁, h₂, Eq.comm (a := a₂)]

/- The following two helper theorems are used to case-split `a = b` representing `iff`. -/
theorem of_eq_eq_true {a b : Prop} (h : (a = b) = True) : (¬a ∨ b) ∧ (¬b ∨ a) := by
by_cases a <;> by_cases b <;> simp_all
theorem of_eq_eq_false {a b : Prop} (h : (a = b) = False) : (¬a ∨ ¬b) ∧ (b ∨ a) := by
by_cases a <;> by_cases b <;> simp_all

/-! Forall -/

theorem forall_propagator (p : Prop) (q : p → Prop) (q' : Prop) (h₁ : p = True) (h₂ : q (of_eq_true h₁) = q') : (∀ hp : p, q hp) = q' := by
Expand Down
22 changes: 15 additions & 7 deletions src/Lean/Meta/Tactic/Grind/Core.lean
Original file line number Diff line number Diff line change
Expand Up @@ -217,14 +217,22 @@ def add (fact : Expr) (proof : Expr) (generation := 0) : GoalM Unit := do
where
go (p : Expr) (isNeg : Bool) : GoalM Unit := do
match_expr p with
| Eq _ lhs rhs => goEq p lhs rhs isNeg false
| HEq _ lhs _ rhs => goEq p lhs rhs isNeg true
| _ =>
internalize p generation
if isNeg then
addEq p (← getFalseExpr) (← mkEqFalse proof)
| Eq α lhs rhs =>
if α.isProp then
-- It is morally an iff.
-- We do not use the `goEq` optimization because we want to register `p` as a case-split
goFact p isNeg
else
addEq p (← getTrueExpr) (← mkEqTrue proof)
goEq p lhs rhs isNeg false
| HEq _ lhs _ rhs => goEq p lhs rhs isNeg true
| _ => goFact p isNeg

goFact (p : Expr) (isNeg : Bool) : GoalM Unit := do
internalize p generation
if isNeg then
addEq p (← getFalseExpr) (← mkEqFalse proof)
else
addEq p (← getTrueExpr) (← mkEqTrue proof)

goEq (p : Expr) (lhs rhs : Expr) (isNeg : Bool) (isHEq : Bool) : GoalM Unit := do
if isNeg then
Expand Down
8 changes: 8 additions & 0 deletions src/Lean/Meta/Tactic/Grind/Internalize.lean
Original file line number Diff line number Diff line change
Expand Up @@ -53,12 +53,20 @@ private def addSplitCandidate (e : Expr) : GoalM Unit := do
-- TODO: add attribute to make this extensible
private def forbiddenSplitTypes := [``Eq, ``HEq, ``True, ``False]

/-- Returns `true` if `e` is of the form `@Eq Prop a b` -/
def isMorallyIff (e : Expr) : Bool :=
let_expr Eq α _ _ := e | false
α.isProp

/-- Inserts `e` into the list of case-split candidates if applicable. -/
private def checkAndAddSplitCandidate (e : Expr) : GoalM Unit := do
unless e.isApp do return ()
if (← getConfig).splitIte && (e.isIte || e.isDIte) then
addSplitCandidate e
return ()
if isMorallyIff e then
addSplitCandidate e
return ()
if (← getConfig).splitMatch then
if (← isMatcherApp e) then
if let .reduced _ ← reduceMatcher? e then
Expand Down
10 changes: 10 additions & 0 deletions src/Lean/Meta/Tactic/Grind/Split.lean
Original file line number Diff line number Diff line change
Expand Up @@ -39,6 +39,11 @@ private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do
return .ready
else
return .notReady
| Eq _ _ _ =>
if (← isEqTrue e <||> isEqFalse e) then
return .ready
else
return .notReady
| ite _ c _ _ _ =>
if (← isEqTrue c <||> isEqFalse c) then
return .resolved
Expand Down Expand Up @@ -94,6 +99,11 @@ private def mkCasesMajor (c : Expr) : GoalM Expr := do
| And a b => return mkApp3 (mkConst ``Grind.or_of_and_eq_false) a b (← mkEqFalseProof c)
| ite _ c _ _ _ => return mkEM c
| dite _ c _ _ _ => return mkEM c
| Eq _ a b =>
if (← isEqTrue c) then
return mkApp3 (mkConst ``Grind.of_eq_eq_true) a b (← mkEqTrueProof c)
else
return mkApp3 (mkConst ``Grind.of_eq_eq_false) a b (← mkEqFalseProof c)
| _ => return mkApp2 (mkConst ``of_eq_true) c (← mkEqTrueProof c)

/-- Introduces new hypotheses in each goal. -/
Expand Down
20 changes: 16 additions & 4 deletions tests/lean/run/grind_implies.lean
Original file line number Diff line number Diff line change
Expand Up @@ -28,10 +28,13 @@ example (p q : Prop) : (p → q) → ¬q → ¬p := by
grind

/--
info: [grind.internalize] p → q
info: [grind.internalize] (p → q) = r
[grind.internalize] Prop
[grind.internalize] p → q
[grind.internalize] p
[grind.internalize] q
[grind.internalize] r
[grind.eqc] ((p → q) = r) = True
[grind.eqc] (p → q) = r
[grind.eqc] p = False
[grind.eqc] (p → q) = True
Expand All @@ -43,10 +46,13 @@ example (p q : Prop) : (p → q) = r → ¬p → r := by


/--
info: [grind.internalize] p → q
info: [grind.internalize] (p → q) = r
[grind.internalize] Prop
[grind.internalize] p → q
[grind.internalize] p
[grind.internalize] q
[grind.internalize] r
[grind.eqc] ((p → q) = r) = True
[grind.eqc] (p → q) = r
[grind.eqc] q = True
[grind.eqc] (p → q) = True
Expand All @@ -57,10 +63,13 @@ example (p q : Prop) : (p → q) = r → q → r := by
grind

/--
info: [grind.internalize] p → q
info: [grind.internalize] (p → q) = r
[grind.internalize] Prop
[grind.internalize] p → q
[grind.internalize] p
[grind.internalize] q
[grind.internalize] r
[grind.eqc] ((p → q) = r) = True
[grind.eqc] (p → q) = r
[grind.eqc] q = False
[grind.eqc] r = True
Expand All @@ -72,10 +81,13 @@ example (p q : Prop) : (p → q) = r → ¬q → r → ¬p := by
grind

/--
info: [grind.internalize] p → q
info: [grind.internalize] (p → q) = r
[grind.internalize] Prop
[grind.internalize] p → q
[grind.internalize] p
[grind.internalize] q
[grind.internalize] r
[grind.eqc] ((p → q) = r) = True
[grind.eqc] (p → q) = r
[grind.eqc] q = False
[grind.eqc] r = False
Expand Down
17 changes: 16 additions & 1 deletion tests/lean/run/grind_pre.lean
Original file line number Diff line number Diff line change
Expand Up @@ -86,13 +86,28 @@ example (p : Prop) (a b c : Nat) : p → a = 0 → a = b → h a = h c → a = c
set_option trace.grind.debug.proof true
/--
error: `grind` failed
case grind
case grind.1
α : Type
a : α
p q r : Prop
h₁ : HEq p a
h₂ : HEq q a
h₃ : p = r
left : ¬p ∨ r
right : ¬r ∨ p
h : ¬r
⊢ False
case grind.2
α : Type
a : α
p q r : Prop
h₁ : HEq p a
h₂ : HEq q a
h₃ : p = r
left : ¬p ∨ r
right : ¬r ∨ p
h : p
⊢ False
-/
#guard_msgs (error) in
Expand Down
9 changes: 9 additions & 0 deletions tests/lean/run/grind_t1.lean
Original file line number Diff line number Diff line change
Expand Up @@ -213,3 +213,12 @@ example (P Q R : α → Prop) (h₁ : ∀ x, Q x → P x) (h₂ : ∀ x, R x →

example (w : Nat → Type) (h : ∀ n, Subsingleton (w n)) : True := by
grind

example {P1 P2 : Prop} : (P1 ∧ P2) ↔ (P2 ∧ P1) := by
grind

example {P U V W : Prop} (h : P ↔ (V ↔ W)) (w : ¬ U ↔ V) : ¬ P ↔ (U ↔ W) := by
grind (splits := 6)

example {P Q : Prop} (q : Q) (w : P = (P = ¬ Q)) : False := by
grind

0 comments on commit f94fc31

Please sign in to comment.