Skip to content

Commit 2a1354b

Browse files
authored
chore: add seal to workaround performance issue (#8469)
This PR adds `seal` commands at `grind_ite.lean` to workaround expensive definitionally equality tests in the canonicalizer. The new module system will automatically hide definitions such as `HashMap.insert` and `TreeMap.insert` which are being unfolded by the canonicalizer in this test. This PR also adds a `profileItM` for tracking the time spent in the `grind` canonicalizer.
1 parent a54872f commit 2a1354b

File tree

2 files changed

+15
-2
lines changed

2 files changed

+15
-2
lines changed

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

Lines changed: 10 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -77,7 +77,15 @@ def canonElemCore (parent : Expr) (f : Expr) (i : Nat) (e : Expr) (useIsDefEqBou
7777
let eType ← inferType e
7878
let cs := s.argMap.find? key |>.getD []
7979
for (c, cType) in cs do
80-
-- We first check the typesr
80+
/-
81+
We first check the types
82+
The following checks are a performance bottleneck.
83+
For example, in the test `grind_ite.lean`, there are many checks of the form:
84+
```
85+
w_4 ∈ assign.insert v true → Prop =?= w_1 ∈ assign.insert v false → Prop
86+
```
87+
where `grind` unfolds the definition of `DHashMap.insert` and `TreeMap.insert`.
88+
-/
8189
if (← withDefault <| isDefEq eType cType) then
8290
if (← isDefEq e c) then
8391
-- We used to check `c.fvarsSubset e` because it is not
@@ -196,7 +204,7 @@ where
196204
end Canon
197205

198206
/-- Canonicalizes nested types, type formers, and instances in `e`. -/
199-
def canon (e : Expr) : GoalM Expr := do
207+
def canon (e : Expr) : GoalM Expr := do profileitM Exception "grind canon" (← getOptions) do
200208
trace_goal[grind.debug.canon] "{e}"
201209
unsafe Canon.canonImpl e
202210

tests/lean/run/grind_ite.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -126,6 +126,11 @@ we are allowed to increase the size of the branches by one, and still be smaller
126126
| var _ => 1
127127
| .ite i t e => 2 * normSize i + max (normSize t) (normSize e) + 1
128128

129+
-- TODO: `grind` canonicalizer is spending a lot of time unfolding the following function.
130+
-- TODO: remove after the new module system will hide this declaration.
131+
seal Std.DHashMap.insert
132+
seal Std.TreeMap.insert
133+
129134
def normalize (assign : Std.HashMap Nat Bool) : IfExpr → IfExpr
130135
| lit b => lit b
131136
| var v =>

0 commit comments

Comments
 (0)