Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

chore: simp tracing reports ← #2621

Merged
merged 5 commits into from
Oct 15, 2023
Merged
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
11 changes: 7 additions & 4 deletions src/Lean/Elab/Tactic/Simp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -138,7 +138,7 @@ def elabSimpArgs (stx : Syntax) (ctx : Simp.Context) (eraseLocal : Bool) (kind :
/-
syntax simpPre := "↓"
syntax simpPost := "↑"
syntax simpLemma := (simpPre <|> simpPost)? term
syntax simpLemma := (simpPre <|> simpPost)? "← "? term

syntax simpErase := "-" ident
-/
Expand Down Expand Up @@ -259,9 +259,12 @@ def traceSimpCall (stx : Syntax) (usedSimps : UsedSimps) : MetaM Unit := do
let env ← getEnv
for (thm, _) in usedSimps.toArray.qsort (·.2 < ·.2) do
match thm with
| .decl declName => -- global definitions in the environment
if env.contains declName && !simpOnlyBuiltins.contains declName then
args := args.push (← `(Parser.Tactic.simpLemma| $(mkIdent (← unresolveNameGlobal declName)):ident))
| .decl declName inv => -- global definitions in the environment
if env.contains declName && (inv || !simpOnlyBuiltins.contains declName) then
args := args.push (if inv then
(← `(Parser.Tactic.simpLemma| ← $(mkIdent (← unresolveNameGlobal declName)):ident))
else
(← `(Parser.Tactic.simpLemma| $(mkIdent (← unresolveNameGlobal declName)):ident)))
| .fvar fvarId => -- local hypotheses in the context
if let some ldecl := lctx.find? fvarId then
localsOrStar := localsOrStar.bind fun locals =>
Expand Down
10 changes: 5 additions & 5 deletions src/Lean/Meta/Tactic/Simp/SimpTheorems.lean
Original file line number Diff line number Diff line change
Expand Up @@ -18,7 +18,7 @@ what action the user took which lead to this theorem existing in the simp set.
-/
inductive Origin where
/-- A global declaration in the environment. -/
| decl (declName : Name)
| decl (declName : Name) (inv := false)
/--
A local hypothesis.
When `contextual := true` is enabled, this fvar may exist in an extension
Expand All @@ -42,7 +42,7 @@ inductive Origin where

/-- A unique identifier corresponding to the origin. -/
def Origin.key : Origin → Name
| .decl declName => declName
| .decl declName _ => declName
| .fvar fvarId => fvarId.name
| .stx id _ => id
| .other name => name
Expand Down Expand Up @@ -136,7 +136,7 @@ instance : ToFormat SimpTheorem where
name ++ prio ++ perm

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

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

Expand Down
8 changes: 8 additions & 0 deletions tests/lean/simp_trace.lean
Original file line number Diff line number Diff line change
Expand Up @@ -126,3 +126,11 @@ instance : HasProp Nat where

example : HasProp.toProp 0 := by
simp [HasProp.toProp]

example (P Q : Prop) (h : P ↔ Q) (p : P) : Q := by
simp [← h]
exact p

theorem my_thm' : a ↔ a ∧ a := my_thm.symm

example (P : Prop) : P ∧ P ↔ P := by simp only [← my_thm']
5 changes: 5 additions & 0 deletions tests/lean/simp_trace.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -115,3 +115,8 @@ Try this: simp only [bla, h] at *
| Sum.inr val => 0
[Meta.Tactic.simp.rewrite] unfold h, h x ==> Sum.inl (x, x)
Try this: simp only [HasProp.toProp]
Try this: simp only [← h]
[Meta.Tactic.simp.rewrite] ← h:1000, Q ==> P
Try this: simp only [← my_thm']
[Meta.Tactic.simp.rewrite] ← @my_thm':1000, P ∧ P ==> P
[Meta.Tactic.simp.rewrite] iff_self:1000, P ↔ P ==> True