Skip to content

Commit 8eee5ff

Browse files
authored
fix: do not include internal match equational theorems at simp trace (#4274)
closes #4251
1 parent fe17b82 commit 8eee5ff

File tree

3 files changed

+28
-3
lines changed

3 files changed

+28
-3
lines changed

src/Lean/Elab/Tactic/Simp.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -336,7 +336,9 @@ def mkSimpOnly (stx : Syntax) (usedSimps : Simp.UsedSimps) : MetaM Syntax := do
336336
for (thm, _) in usedSimps.toArray.qsort (·.2 < ·.2) do
337337
match thm with
338338
| .decl declName post inv => -- global definitions in the environment
339-
if env.contains declName && (inv || !simpOnlyBuiltins.contains declName) then
339+
if env.contains declName
340+
&& (inv || !simpOnlyBuiltins.contains declName)
341+
&& !Match.isMatchEqnTheorem env declName then
340342
let decl : Term ← `($(mkIdent (← unresolveNameGlobal declName)):ident)
341343
let arg ← match post, inv with
342344
| true, true => `(Parser.Tactic.simpLemma| ← $decl:term)

src/Lean/Meta/Match/MatchEqsExt.lean

Lines changed: 12 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -19,19 +19,29 @@ def MatchEqns.size (e : MatchEqns) : Nat :=
1919

2020
structure MatchEqnsExtState where
2121
map : PHashMap Name MatchEqns := {}
22+
eqns : PHashSet Name := {}
2223
deriving Inhabited
2324

2425
/- We generate the equations and splitter on demand, and do not save them on .olean files. -/
2526
builtin_initialize matchEqnsExt : EnvExtension MatchEqnsExtState ←
2627
registerEnvExtension (pure {})
2728

28-
def registerMatchEqns (matchDeclName : Name) (matchEqns : MatchEqns) : CoreM Unit :=
29-
modifyEnv fun env => matchEqnsExt.modifyState env fun s => { s with map := s.map.insert matchDeclName matchEqns }
29+
def registerMatchEqns (matchDeclName : Name) (matchEqns : MatchEqns) : CoreM Unit := do
30+
modifyEnv fun env => matchEqnsExt.modifyState env fun { map, eqns } => {
31+
eqns := matchEqns.eqnNames.foldl (init := eqns) fun eqns eqn => eqns.insert eqn
32+
map := map.insert matchDeclName matchEqns
33+
}
3034

3135
/-
3236
Forward definition. We want to use `getEquationsFor` in the simplifier,
3337
`getEquationsFor` depends on `mkEquationsfor` which uses the simplifier. -/
3438
@[extern "lean_get_match_equations_for"]
3539
opaque getEquationsFor (matchDeclName : Name) : MetaM MatchEqns
3640

41+
/--
42+
Returns `true` if `declName` is the name of a `match` equational theorem.
43+
-/
44+
def isMatchEqnTheorem (env : Environment) (declName : Name) : Bool :=
45+
matchEqnsExt.getState env |>.eqns.contains declName
46+
3747
end Lean.Meta.Match

tests/lean/run/4251.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,13 @@
1+
/--
2+
info: Try this: simp only [ha, Nat.reduceEqDiff, imp_self]
3+
-/
4+
#guard_msgs in
5+
theorem foo₁ (a : Nat) (ha : a = 37) :
6+
(match h : a with | 42 => by simp_all | n => n) = 37 := by
7+
simp? [ha]
8+
9+
theorem foo₂ (a : Nat) (ha : a = 37) (hb : b = 37) :
10+
(match h : a with | 42 => by simp_all | n => n) = b := by
11+
simp only [ha, Nat.reduceEqDiff, imp_self]
12+
guard_target =ₛ 37 = b
13+
rw [hb]

0 commit comments

Comments
 (0)