@@ -269,19 +269,36 @@ private def getPatternFn? (pattern : Expr) : Option Expr :=
269
269
270
270
/--
271
271
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
273
273
- a proof, or
274
274
- an instance implicit argument
275
275
276
276
When `mask[i]`, we say the corresponding argument is a "support" argument.
277
277
-/
278
278
def getPatternSupportMask (f : Expr) (numArgs : Nat) : MetaM (Array Bool) := do
279
+ let pinfos := (← getFunInfoNArgs f numArgs).paramInfo
279
280
forallBoundedTelescope (← inferType f) numArgs fun xs _ => do
280
- xs.mapM fun x => do
281
+ xs.mapIdxM fun idx x => do
281
282
if (← isProp x) then
282
283
return false
283
- else if (← isTypeFormer x <||> isProof x) then
284
+ else if (← isProof x) then
284
285
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
285
302
else
286
303
return (← x.fvarId!.getDecl).binderInfo matches .instImplicit
287
304
0 commit comments