Skip to content

Commit cd50e67

Browse files
committed
fix: preprocessLight at ensureInternalized
1 parent 2b0b1e0 commit cd50e67

File tree

2 files changed

+7
-9
lines changed

2 files changed

+7
-9
lines changed

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

Lines changed: 0 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -17,14 +17,6 @@ import Lean.Meta.Tactic.Grind.MatchCond
1717
import Lean.Meta.Tactic.Grind.Arith.Internalize
1818

1919
namespace Lean.Meta.Grind
20-
21-
/--
22-
A lighter version of `preprocess` which produces a definitionally equal term,
23-
but ensures assumptions made by `grind` are satisfied.
24-
-/
25-
private def preprocessLight (e : Expr) : GoalM Expr := do
26-
shareCommon (← canon (← normalizeLevels (← foldProjs (← eraseIrrelevantMData (← markNestedProofs (← unfoldReducible e))))))
27-
2820
/-- Adds `e` to congruence table. -/
2921
def addCongrTable (e : Expr) : GoalM Unit := do
3022
if let some { e := e' } := (← get).congrTable.find? { e } then

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

Lines changed: 7 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,12 @@ import Lean.Meta.Tactic.Grind.Types
88
import Lean.Meta.Tactic.Grind.Simp
99

1010
namespace Lean.Meta.Grind
11+
/--
12+
A lighter version of `preprocess` which produces a definitionally equal term,
13+
but ensures assumptions made by `grind` are satisfied.
14+
-/
15+
def preprocessLight (e : Expr) : GoalM Expr := do
16+
shareCommon (← canon (← normalizeLevels (← foldProjs (← eraseIrrelevantMData (← markNestedProofs (← unfoldReducible e))))))
1117

1218
/--
1319
If `e` has not been internalized yet, instantiate metavariables, unfold reducible, canonicalize,
@@ -24,7 +30,7 @@ private def ensureInternalized (e : Expr) : GoalM Expr := do
2430
`¬ a = []` and `b ≠ []` by congruence closure even when `a` and `b` are in the same
2531
equivalence class.
2632
-/
27-
let e ← shareCommon (← canon (← unfoldReducible (← instantiateMVars e)))
33+
let e ← preprocessLight (← instantiateMVars e)
2834
internalize e 0
2935
return e
3036

0 commit comments

Comments
 (0)