Skip to content

Commit 4eba5ea

Browse files
authored
fix: shake: only record non-builtin simprocs (#11344)
1 parent 075f1d6 commit 4eba5ea

File tree

2 files changed

+17
-8
lines changed

2 files changed

+17
-8
lines changed

src/Lean/Meta/Tactic/Simp/Main.lean

Lines changed: 17 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -12,6 +12,7 @@ public import Lean.Meta.Tactic.Simp.Diagnostics
1212
public import Lean.Meta.Match.Value
1313
public import Lean.Util.CollectLooseBVars
1414
import Lean.PrettyPrinter
15+
import Lean.ExtraModUses
1516

1617
public section
1718

@@ -1069,15 +1070,29 @@ def simpImpl (e : Expr) : SimpM Result := withIncRecDepth do
10691070
else
10701071
x
10711072

1073+
/--
1074+
For `rfl` theorems and simprocs, there might not be an explicit reference in the proof term, so
1075+
we record (all non-builtin) usages explicitly.
1076+
-/
1077+
private def recordSimpUses (s : State) : MetaM Unit := do
1078+
for (thm, _) in s.usedTheorems.map do
1079+
if let .decl declName .. := thm then
1080+
if !(← isBuiltinSimproc declName) then
1081+
recordExtraModUseFromDecl (isMeta := false) declName
1082+
10721083
def mainCore (e : Expr) (ctx : Context) (s : State := {}) (methods : Methods := {}) : MetaM (Result × State) := do
1073-
SimpM.run ctx s methods <| withCatchingRuntimeEx <| simp e
1084+
let (r, s) ← SimpM.run ctx s methods <| withCatchingRuntimeEx <| simp e
1085+
recordSimpUses s
1086+
return (r, s)
10741087

10751088
def main (e : Expr) (ctx : Context) (stats : Stats := {}) (methods : Methods := {}) : MetaM (Result × Stats) := do
10761089
let (r, s) ← mainCore e ctx { stats with } methods
10771090
return (r, { s with })
10781091

10791092
def dsimpMainCore (e : Expr) (ctx : Context) (s : State := {}) (methods : Methods := {}) : MetaM (Expr × State) := do
1080-
SimpM.run ctx s methods <| withCatchingRuntimeEx <| dsimp e
1093+
let (r, s) ← SimpM.run ctx s methods <| withCatchingRuntimeEx <| dsimp e
1094+
recordSimpUses s
1095+
return (r, s)
10811096

10821097
def dsimpMain (e : Expr) (ctx : Context) (stats : Stats := {}) (methods : Methods := {}) : MetaM (Expr × Stats) := do
10831098
let (r, s) ← dsimpMainCore e ctx { stats with } methods

src/Lean/Meta/Tactic/Simp/Types.lean

Lines changed: 0 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -12,7 +12,6 @@ public import Lean.Meta.Eqns
1212
public import Lean.Meta.Tactic.Simp.SimpTheorems
1313
public import Lean.Meta.Tactic.Simp.SimpCongrTheorems
1414
import Lean.Meta.Tactic.Replace
15-
import Lean.ExtraModUses
1615

1716
public section
1817

@@ -881,11 +880,6 @@ def SimpM.run (ctx : Context) (s : State := {}) (methods : Methods := {}) (k : S
881880
trace[Meta.Tactic.simp.numSteps] "{s.numSteps}"
882881
let stats ← updateUsedSimpsWithZetaDelta ctx { s with }
883882
let s := { s with diag := stats.diag, usedTheorems := stats.usedTheorems }
884-
for (thm, _) in s.usedTheorems.map do
885-
if let .decl declName .. := thm then
886-
-- for `rfl` theorems and simprocs, there might not be an explicit reference in the proof
887-
-- term
888-
recordExtraModUseFromDecl (isMeta := false) declName
889883
return (r, s)
890884

891885
end Simp

0 commit comments

Comments
 (0)