Skip to content

Commit 5f2f010

Browse files
authored
fix: missing forall normalization rules in grind (#7808)
This PR adds missing forall normalization rules to `grind`.
1 parent 29303b3 commit 5f2f010

File tree

3 files changed

+14
-10
lines changed

3 files changed

+14
-10
lines changed

src/Init/Grind/Norm.lean

Lines changed: 5 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -83,6 +83,9 @@ theorem Nat.pow_one (a : Nat) : a ^ 1 = a := by
8383
theorem Int.pow_one (a : Int) : a ^ 1 = a := by
8484
simp [Int.pow_succ]
8585

86+
theorem forall_true (p : True → Prop) : (∀ h : True, p h) = p True.intro :=
87+
propext <| Iff.intro (fun h => h True.intro) (fun h _ => h)
88+
8689
init_grind_norm
8790
/- Pre theorems -/
8891
not_and not_or not_ite not_forall not_exists
@@ -108,7 +111,7 @@ init_grind_norm
108111
ite_true ite_false ite_true_false ite_false_true
109112
dite_eq_ite
110113
-- Forall
111-
forall_and
114+
forall_and forall_false forall_true
112115
-- Exists
113116
exists_const exists_or exists_prop exists_and_left exists_and_right
114117
-- Bool cond
@@ -131,7 +134,7 @@ init_grind_norm
131134
Nat.le_zero_eq Nat.lt_eq Nat.succ_eq_add_one
132135
Nat.add_eq Nat.sub_eq Nat.mul_eq Nat.zero_eq Nat.le_eq
133136
Nat.div_zero Nat.mod_zero Nat.div_one Nat.mod_one
134-
Nat.sub_sub Nat.pow_zero Nat.pow_one
137+
Nat.sub_sub Nat.pow_zero Nat.pow_one Nat.sub_self
135138
-- Int
136139
Int.lt_eq
137140
Int.emod_neg Int.ediv_neg

tests/lean/grind/list_problems.lean

Lines changed: 5 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -2,18 +2,15 @@ theorem getElem?_eq_some_iff {l : List α} : l[i]? = some a ↔ ∃ h : i < l.le
22
induction l
33
· grind
44
· cases i
5-
· -- Fails because of the issue:
6-
-- [issue] failed to create E-match local theorem for
7-
-- ∀ (x : 1 ≤ tail.length), ¬tail[0] = a
8-
-- despite having asserted `1 ≤ tail.length `.
5+
· -- Better support for implication and dependent implication.
6+
-- We need inequality propagation (or case-splits)
97
grind
108
· -- Similarly
119
grind
1210

1311
attribute [grind] List.getElem_append_left List.getElem_append_right
12+
attribute [grind] List.length_cons List.length_nil
1413

15-
@[simp] theorem getElem_concat_length {l : List α} {a : α} {i : Nat} (h : i = l.length) (w) :
14+
example {l : List α} {a : α} {i : Nat} (h : i = l.length) (w) :
1615
(l ++ [a])[i]'w = a := by
17-
subst h; grind
18-
-- [issue] failed to create E-match local theorem for
19-
-- ∀ (h₁ : True), (l ++ [a])[l.length] = [a][l.length - l.length]
16+
grind -- Similar to issue above.

tests/lean/run/grind_t1.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -420,3 +420,7 @@ example [BEq α] [LawfulBEq α] (a b : α) : a ≠ b → foo a b = 0 := by
420420

421421
example [BEq α] [LawfulBEq α] (a b : α) : a ≠ b → foo a b = 0 := by
422422
grind (splits := 0) [foo]
423+
424+
@[simp] theorem getElem_concat_length {l : List α} {a : α} {i : Nat} (h : i = l.length) (w) :
425+
(l ++ [a])[i]'w = a := by
426+
subst h; grind [List.getElem_append_left, List.getElem_append_right]

0 commit comments

Comments
 (0)