Skip to content

Commit 2158c9e

Browse files
committed
fix: generalize excessive resource usage
closes #3524
1 parent 870de43 commit 2158c9e

File tree

2 files changed

+64
-10
lines changed

2 files changed

+64
-10
lines changed

src/Lean/Meta/Tactic/Generalize.lean

+61-10
Original file line numberDiff line numberDiff line change
@@ -23,7 +23,8 @@ structure GeneralizeArg where
2323
Telescopic `generalize` tactic. It can simultaneously generalize many terms.
2424
It uses `kabstract` to occurrences of the terms that need to be generalized.
2525
-/
26-
private partial def generalizeCore (mvarId : MVarId) (args : Array GeneralizeArg) : MetaM (Array FVarId × MVarId) :=
26+
private partial def generalizeCore (mvarId : MVarId) (args : Array GeneralizeArg)
27+
(transparency : TransparencyMode) : MetaM (Array FVarId × MVarId) :=
2728
mvarId.withContext do
2829
mvarId.checkNotAssigned `generalize
2930
let tag ← mvarId.getTag
@@ -35,7 +36,8 @@ private partial def generalizeCore (mvarId : MVarId) (args : Array GeneralizeArg
3536
let eType ← instantiateMVars (← inferType e)
3637
let type ← go (i+1)
3738
let xName ← if let some xName := arg.xName? then pure xName else mkFreshUserName `x
38-
return Lean.mkForall xName BinderInfo.default eType (← kabstract type e)
39+
return Lean.mkForall xName BinderInfo.default eType
40+
(← withTransparency transparency <| kabstract type e)
3941
else
4042
return target
4143
let targetNew ← go 0
@@ -71,30 +73,79 @@ private partial def generalizeCore (mvarId : MVarId) (args : Array GeneralizeArg
7173
mvarId.assign (mkAppN (mkAppN mvarNew es) rfls.toArray)
7274
mvarNew.mvarId!.introNP (args.size + rfls.length)
7375

76+
/-
77+
Remark: we use `TransparencyMode.instances` as the default setting at `generalize`
78+
and `generalizeHyp` to avoid excessive resource usage.
79+
80+
**Motivation:**
81+
The `kabstract e p` operation is widely used, for instance, in the `generalize` tactic.
82+
It operates by taking an expression `e` and a pattern (i.e., an expression containing metavariables)
83+
and employs keyed-matching to identify and abstract instances of `p` within `e`.
84+
For example, if `e` is `a + (2 * (b + c))` and `p` is `2 * ?m`, the resultant expression
85+
would be `a + #0`, where `#0` represents a loose bound variable.
86+
87+
This matching process is not merely syntactic; it also considers reduction. It's impractical
88+
to attempt matching each sub-term with `p`; therefore, only sub-terms sharing the same "root"
89+
symbol are evaluated. For instance, with the pattern `2 * ?m`, only sub-terms with the
90+
root `*` are considered. Matching is executed using the definitionally equality test
91+
(i.e., `isDefEq`).
92+
93+
The `generalize` tactic employs `kabstract` and defaults to standard reducibility.
94+
Hence, the `isDefEq` operations invoked by `kabstract` can become highly resource-intensive
95+
and potentially trigger "max recursion depth reached" errors, as observed in issue #3524.
96+
This issue was isolated by @**Scott Morrison** with the following example:
97+
```
98+
example (a : Nat) : ((2 ^ 7) + a) - 2 ^ 7 = 0 := by
99+
generalize 0 - 0 = x
100+
```
101+
In this scenario, `kabstract` triggers a "max recursion depth reached" error while
102+
testing whether `((2 ^ 7) + a) - 2 ^ 7` is definitionally equal to `0 - 0`.
103+
Note that the term `((2 ^ 7) + a) - 2 ^ 7` is not ground.
104+
We believe most users find the error message to be uninformative and unexpected.
105+
To fix this issue, we decided to use `TransparencyMode.instances` as the default setting.
106+
107+
Kyle Miller has performed the following analysis on the potential impact of the
108+
changes on Mathlib (2024-03-02).
109+
110+
There seem to be just 130 cases of generalize in Mathlib, and after looking through a
111+
good number of them, they seem to come in just two types:
112+
113+
- Ones where it looks like reducible+instance transparency should work, where in
114+
particular there is nothing obvious being reduced, and
115+
- Ones that don't make use of the `kabstract` feature at all (it's being used like a
116+
`have` that introduces an equality for rewriting).
117+
118+
That wasn't a systematic review of generalize though. It's possible changing the
119+
transparency settings would break things, but in my opinion it would be better
120+
if generalize weren't used for unfolding things.
121+
-/
122+
74123
@[inherit_doc generalizeCore]
75-
def _root_.Lean.MVarId.generalize (mvarId : MVarId) (args : Array GeneralizeArg) : MetaM (Array FVarId × MVarId) :=
76-
generalizeCore mvarId args
124+
def _root_.Lean.MVarId.generalize (mvarId : MVarId) (args : Array GeneralizeArg)
125+
(transparency := TransparencyMode.instances) : MetaM (Array FVarId × MVarId) :=
126+
generalizeCore mvarId args transparency
77127

78128
@[inherit_doc generalizeCore, deprecated MVarId.generalize]
79-
def generalize (mvarId : MVarId) (args : Array GeneralizeArg) : MetaM (Array FVarId × MVarId) :=
80-
generalizeCore mvarId args
129+
def generalize (mvarId : MVarId) (args : Array GeneralizeArg)
130+
(transparency := TransparencyMode.instances) : MetaM (Array FVarId × MVarId) :=
131+
generalizeCore mvarId args transparency
81132

82133
/--
83134
Extension of `generalize` to support generalizing within specified hypotheses.
84135
The `hyps` array contains the list of hypotheses within which to look for occurrences
85136
of the generalizing expressions.
86137
-/
87138
def _root_.Lean.MVarId.generalizeHyp (mvarId : MVarId) (args : Array GeneralizeArg) (hyps : Array FVarId := #[])
88-
(fvarSubst : FVarSubst := {}) : MetaM (FVarSubst × Array FVarId × MVarId) := do
139+
(fvarSubst : FVarSubst := {}) (transparency := TransparencyMode.instances) : MetaM (FVarSubst × Array FVarId × MVarId) := do
89140
if hyps.isEmpty then
90141
-- trivial case
91-
return (fvarSubst, ← mvarId.generalize args)
142+
return (fvarSubst, ← mvarId.generalize args transparency)
92143
let args ← args.mapM fun arg => return { arg with expr := ← instantiateMVars arg.expr }
93144
let hyps ← hyps.filterM fun h => do
94145
let type ← instantiateMVars (← h.getType)
95-
args.anyM fun arg => return (← kabstract type arg.expr).hasLooseBVars
146+
args.anyM fun arg => return (← withTransparency transparency <| kabstract type arg.expr).hasLooseBVars
96147
let (reverted, mvarId) ← mvarId.revert hyps true
97-
let (newVars, mvarId) ← mvarId.generalize args
148+
let (newVars, mvarId) ← mvarId.generalize args transparency
98149
let (reintros, mvarId) ← mvarId.introNP reverted.size
99150
let fvarSubst := Id.run do
100151
let mut subst : FVarSubst := fvarSubst

tests/lean/run/3524.lean

+3
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,3 @@
1+
example (a : Nat) : ((2 ^ 7) + a) - 2 ^ 7 = 0 := by
2+
generalize 0 - 0 = x
3+
sorry

0 commit comments

Comments
 (0)