Skip to content

Commit b3f8fef

Browse files
authored
fix: improve E-matching pattern selection heuristics (#6653)
This PR improves the E-matching pattern selection heuristics in the `grind` tactic. They now take into account type predicates and transformers.
1 parent 6665837 commit b3f8fef

File tree

2 files changed

+23
-3
lines changed

2 files changed

+23
-3
lines changed

src/Lean/Meta/Tactic/Grind/EMatchTheorem.lean

Lines changed: 20 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -269,19 +269,36 @@ private def getPatternFn? (pattern : Expr) : Option Expr :=
269269

270270
/--
271271
Returns a bit-mask `mask` s.t. `mask[i]` is true if the corresponding argument is
272-
- a type (that is not a proposition) or type former, or
272+
- a type (that is not a proposition) or type former (which has forward dependencies) or
273273
- a proof, or
274274
- an instance implicit argument
275275
276276
When `mask[i]`, we say the corresponding argument is a "support" argument.
277277
-/
278278
def getPatternSupportMask (f : Expr) (numArgs : Nat) : MetaM (Array Bool) := do
279+
let pinfos := (← getFunInfoNArgs f numArgs).paramInfo
279280
forallBoundedTelescope (← inferType f) numArgs fun xs _ => do
280-
xs.mapM fun x => do
281+
xs.mapIdxM fun idx x => do
281282
if (← isProp x) then
282283
return false
283-
else if (← isTypeFormer x <||> isProof x) then
284+
else if (← isProof x) then
284285
return true
286+
else if (← isTypeFormer x) then
287+
if h : idx < pinfos.size then
288+
/-
289+
We originally wanted to ignore types and type formers in `grind` and treat them as supporting elements.
290+
Thus, we would always return `true`. However, we changed our heuristic because of the following example:
291+
```
292+
example {α} (f : α → Type) (a : α) (h : ∀ x, Nonempty (f x)) : Nonempty (f a) := by
293+
grind
294+
```
295+
In this example, we are reasoning about types. Therefore, we adjusted the heuristic as follows:
296+
a type or type former is considered a supporting element only if it has forward dependencies.
297+
Note that this is not the case for `Nonempty`.
298+
-/
299+
return pinfos[idx].hasFwdDeps
300+
else
301+
return true
285302
else
286303
return (← x.fvarId!.getDecl).binderInfo matches .instImplicit
287304

tests/lean/run/grind_t1.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -280,3 +280,6 @@ example {a b : Nat} (h : a < b) : ¬ b < a := by
280280

281281
example {m n : Nat} : m < n ↔ m ≤ n ∧ ¬ n ≤ m := by
282282
grind
283+
284+
example {α} (f : α → Type) (a : α) (h : ∀ x, Nonempty (f x)) : Nonempty (f a) := by
285+
grind

0 commit comments

Comments
 (0)