Skip to content

Commit 66ab016

Browse files
kim-emdigama0
andauthored
chore: simp tracing reports ← (#2621)
* chore: simp tracing reports ← --------- Co-authored-by: Mario Carneiro <[email protected]>
1 parent 6df09d1 commit 66ab016

File tree

4 files changed

+25
-9
lines changed

4 files changed

+25
-9
lines changed

src/Lean/Elab/Tactic/Simp.lean

+7-4
Original file line numberDiff line numberDiff line change
@@ -138,7 +138,7 @@ def elabSimpArgs (stx : Syntax) (ctx : Simp.Context) (eraseLocal : Bool) (kind :
138138
/-
139139
syntax simpPre := "↓"
140140
syntax simpPost := "↑"
141-
syntax simpLemma := (simpPre <|> simpPost)? term
141+
syntax simpLemma := (simpPre <|> simpPost)? "← "? term
142142
143143
syntax simpErase := "-" ident
144144
-/
@@ -259,9 +259,12 @@ def traceSimpCall (stx : Syntax) (usedSimps : UsedSimps) : MetaM Unit := do
259259
let env ← getEnv
260260
for (thm, _) in usedSimps.toArray.qsort (·.2 < ·.2) do
261261
match thm with
262-
| .decl declName => -- global definitions in the environment
263-
if env.contains declName && !simpOnlyBuiltins.contains declName then
264-
args := args.push (← `(Parser.Tactic.simpLemma| $(mkIdent (← unresolveNameGlobal declName)):ident))
262+
| .decl declName inv => -- global definitions in the environment
263+
if env.contains declName && (inv || !simpOnlyBuiltins.contains declName) then
264+
args := args.push (if inv then
265+
(← `(Parser.Tactic.simpLemma| ← $(mkIdent (← unresolveNameGlobal declName)):ident))
266+
else
267+
(← `(Parser.Tactic.simpLemma| $(mkIdent (← unresolveNameGlobal declName)):ident)))
265268
| .fvar fvarId => -- local hypotheses in the context
266269
if let some ldecl := lctx.find? fvarId then
267270
localsOrStar := localsOrStar.bind fun locals =>

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

+5-5
Original file line numberDiff line numberDiff line change
@@ -18,7 +18,7 @@ what action the user took which lead to this theorem existing in the simp set.
1818
-/
1919
inductive Origin where
2020
/-- A global declaration in the environment. -/
21-
| decl (declName : Name)
21+
| decl (declName : Name) (inv := false)
2222
/--
2323
A local hypothesis.
2424
When `contextual := true` is enabled, this fvar may exist in an extension
@@ -42,7 +42,7 @@ inductive Origin where
4242

4343
/-- A unique identifier corresponding to the origin. -/
4444
def Origin.key : Origin → Name
45-
| .decl declName => declName
45+
| .decl declName _ => declName
4646
| .fvar fvarId => fvarId.name
4747
| .stx id _ => id
4848
| .other name => name
@@ -136,7 +136,7 @@ instance : ToFormat SimpTheorem where
136136
name ++ prio ++ perm
137137

138138
def ppOrigin [Monad m] [MonadEnv m] [MonadError m] : Origin → m MessageData
139-
| .decl n => mkConstWithLevelParams n
139+
| .decl n inv => do let r ← mkConstWithLevelParams n; if inv then return m!"← {r}" else return r
140140
| .fvar n => return mkFVar n
141141
| .stx _ ref => return ref
142142
| .other n => return n
@@ -318,7 +318,7 @@ private def mkSimpTheoremsFromConst (declName : Name) (post : Bool) (inv : Bool)
318318
let mut r := #[]
319319
for (val, type) in (← preprocess val type inv (isGlobal := true)) do
320320
let auxName ← mkAuxLemma cinfo.levelParams type val
321-
r := r.push <| (← mkSimpTheoremCore (.decl declName) (mkConst auxName (cinfo.levelParams.map mkLevelParam)) #[] (mkConst auxName) post prio)
321+
r := r.push <| (← mkSimpTheoremCore (.decl declName inv) (mkConst auxName (cinfo.levelParams.map mkLevelParam)) #[] (mkConst auxName) post prio)
322322
return r
323323
else
324324
return #[← mkSimpTheoremCore (.decl declName) (mkConst declName (cinfo.levelParams.map mkLevelParam)) #[] (mkConst declName) post prio]
@@ -403,7 +403,7 @@ def getSimpTheorems : CoreM SimpTheorems :=
403403

404404
/-- Auxiliary method for adding a global declaration to a `SimpTheorems` datastructure. -/
405405
def SimpTheorems.addConst (s : SimpTheorems) (declName : Name) (post := true) (inv := false) (prio : Nat := eval_prio default) : MetaM SimpTheorems := do
406-
let s := { s with erased := s.erased.erase (.decl declName) }
406+
let s := { s with erased := s.erased.erase (.decl declName inv) }
407407
let simpThms ← mkSimpTheoremsFromConst declName post inv prio
408408
return simpThms.foldl addSimpTheoremEntry s
409409

tests/lean/simp_trace.lean

+8
Original file line numberDiff line numberDiff line change
@@ -126,3 +126,11 @@ instance : HasProp Nat where
126126

127127
example : HasProp.toProp 0 := by
128128
simp [HasProp.toProp]
129+
130+
example (P Q : Prop) (h : P ↔ Q) (p : P) : Q := by
131+
simp [← h]
132+
exact p
133+
134+
theorem my_thm' : a ↔ a ∧ a := my_thm.symm
135+
136+
example (P : Prop) : P ∧ P ↔ P := by simp only [← my_thm']

tests/lean/simp_trace.lean.expected.out

+5
Original file line numberDiff line numberDiff line change
@@ -115,3 +115,8 @@ Try this: simp only [bla, h] at *
115115
| Sum.inr val => 0
116116
[Meta.Tactic.simp.rewrite] unfold h, h x ==> Sum.inl (x, x)
117117
Try this: simp only [HasProp.toProp]
118+
Try this: simp only [← h]
119+
[Meta.Tactic.simp.rewrite] ← h:1000, Q ==> P
120+
Try this: simp only [← my_thm']
121+
[Meta.Tactic.simp.rewrite] ← @my_thm':1000, P ∧ P ==> P
122+
[Meta.Tactic.simp.rewrite] iff_self:1000, P ↔ P ==> True

0 commit comments

Comments
 (0)