Skip to content

Commit

Permalink
feat: add simp? and dsimp? in conversion mode (#6593)
Browse files Browse the repository at this point in the history
This PR adds support for the `simp?` and `dsimp?` tactics in conversion
mode.

Closes #6164
  • Loading branch information
jrr6 authored Jan 10, 2025
1 parent d369976 commit c07948a
Show file tree
Hide file tree
Showing 3 changed files with 62 additions and 1 deletion.
7 changes: 7 additions & 0 deletions src/Init/Conv.lean
Original file line number Diff line number Diff line change
Expand Up @@ -150,6 +150,10 @@ See the `simp` tactic for more information. -/
syntax (name := simp) "simp" optConfig (discharger)? (&" only")?
(" [" withoutPosition((simpStar <|> simpErase <|> simpLemma),*) "]")? : conv

/-- `simp?` takes the same arguments as `simp`, but reports an equivalent call to `simp only`
that would be sufficient to close the goal. See the `simp?` tactic for more information. -/
syntax (name := simpTrace) "simp?" optConfig (discharger)? (&" only")? (simpArgs)? : conv

/--
`dsimp` is the definitional simplifier in `conv`-mode. It differs from `simp` in that it only
applies theorems that hold by reflexivity.
Expand All @@ -167,6 +171,9 @@ example (a : Nat): (0 + 0) = a - a := by
syntax (name := dsimp) "dsimp" optConfig (discharger)? (&" only")?
(" [" withoutPosition((simpErase <|> simpLemma),*) "]")? : conv

@[inherit_doc simpTrace]
syntax (name := dsimpTrace) "dsimp?" optConfig (&" only")? (dsimpArgs)? : conv

/-- `simp_match` simplifies match expressions. For example,
```
match [a, b] with
Expand Down
27 changes: 26 additions & 1 deletion src/Lean/Elab/Tactic/Conv/Simp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -7,9 +7,10 @@ prelude
import Lean.Elab.Tactic.Simp
import Lean.Elab.Tactic.Split
import Lean.Elab.Tactic.Conv.Basic
import Lean.Elab.Tactic.SimpTrace

namespace Lean.Elab.Tactic.Conv
open Meta
open Meta Tactic TryThis

def applySimpResult (result : Simp.Result) : TacticM Unit := do
if result.proof?.isNone then
Expand All @@ -23,11 +24,35 @@ def applySimpResult (result : Simp.Result) : TacticM Unit := do
let (result, _) ← dischargeWrapper.with fun d? => simp lhs ctx (simprocs := simprocs) (discharge? := d?)
applySimpResult result

@[builtin_tactic Lean.Parser.Tactic.Conv.simpTrace] def evalSimpTrace : Tactic := fun stx => withMainContext do
match stx with
| `(conv| simp?%$tk $cfg:optConfig $(discharger)? $[only%$o]? $[[$args,*]]?) => do
let stx ← `(tactic| simp%$tk $cfg:optConfig $[$discharger]? $[only%$o]? $[[$args,*]]?)
let { ctx, simprocs, dischargeWrapper, .. } ← mkSimpContext stx (eraseLocal := false)
let lhs ← getLhs
let (result, stats) ← dischargeWrapper.with fun d? =>
simp lhs ctx (simprocs := simprocs) (discharge? := d?)
applySimpResult result
let stx ← mkSimpCallStx stx stats.usedTheorems
addSuggestion tk stx (origSpan? := ← getRef)
| _ => throwUnsupportedSyntax

@[builtin_tactic Lean.Parser.Tactic.Conv.simpMatch] def evalSimpMatch : Tactic := fun _ => withMainContext do
applySimpResult (← Split.simpMatch (← getLhs))

@[builtin_tactic Lean.Parser.Tactic.Conv.dsimp] def evalDSimp : Tactic := fun stx => withMainContext do
let { ctx, .. } ← mkSimpContext stx (eraseLocal := false) (kind := .dsimp)
changeLhs (← Lean.Meta.dsimp (← getLhs) ctx).1

@[builtin_tactic Lean.Parser.Tactic.Conv.dsimpTrace] def evalDSimpTrace : Tactic := fun stx => withMainContext do
match stx with
| `(conv| dsimp?%$tk $cfg:optConfig $[only%$o]? $[[$args,*]]?) =>
let stx ← `(tactic| dsimp%$tk $cfg:optConfig $[only%$o]? $[[$args,*]]?)
let { ctx, .. } ← mkSimpContext stx (eraseLocal := false) (kind := .dsimp)
let (result, stats) ← Lean.Meta.dsimp (← getLhs) ctx
changeLhs result
let stx ← mkSimpCallStx stx stats.usedTheorems
addSuggestion tk stx (origSpan? := ← getRef)
| _ => throwUnsupportedSyntax

end Lean.Elab.Tactic.Conv
29 changes: 29 additions & 0 deletions tests/lean/run/6164.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,29 @@
/-!
# `simp?` in conversion mode
Tests that `simp?` and `dsimp?` work properly in `conv` mode: namely, that each displays the
appropriate suggestion and applies the corresponding simplification to the focused expression.
-/

attribute [simp] Nat.two_mul

/--
info: Try this: simp only [Nat.two_mul]
-/
#guard_msgs in
example (n : Nat) : 123 + 2 * n = 123 + (n + n) := by
conv =>
enter [1, 2]
simp?


@[simp] def foo (n : Nat) := n + 1

/--
info: Try this: dsimp only [foo]
-/
#guard_msgs in
example (n : Nat) : foo n = n + 1 := by
conv =>
lhs
dsimp?

0 comments on commit c07948a

Please sign in to comment.