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

feat: add Simprocs #3124

Merged
merged 54 commits into from
Jan 9, 2024
Merged
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
54 commits
Select commit Hold shift + click to select a range
2ea2ef5
perf: (try to) fix regression introduced by #3139
leodemoura Jan 7, 2024
7e2f93b
feat: add `reduceStep`, and try `pre` simp steps again if term was re…
leodemoura Dec 26, 2023
979a315
feat: add `Expr.getAppPrefix`
leodemoura Dec 26, 2023
d2a40c8
feat: add `Expr.getAppArgsN`
leodemoura Dec 26, 2023
cac7160
feat: better support for `match`-application in the simplifier
leodemoura Dec 27, 2023
81d4336
feat: add pre simp lemmas for if-then-else terms
leodemoura Dec 27, 2023
ec958f3
chore: fix regression due to changes in previous commits
leodemoura Dec 27, 2023
7b2dc2a
refactor: use `unsafe` code to break recursion in `simp` implementation
leodemoura Dec 27, 2023
0041d4d
chore: simplify `mutual` at `simpImpl`
leodemoura Dec 27, 2023
102928b
refactor: simplify `match`-expressions at `pre` simp method
leodemoura Dec 27, 2023
8aac0fe
refactor: simplify `simpImpl`
leodemoura Dec 27, 2023
0055018
chore: address feedback
leodemoura Dec 30, 2023
312b8eb
feat: add `simproc`s
leodemoura Dec 28, 2023
3060997
feat: simproc declaration vs simproc attribute
leodemoura Dec 29, 2023
9f09264
feat: `simp only` should not use default `simproc` set
leodemoura Dec 29, 2023
6e9ea19
feat: allow extra `simproc`s to be provided as `simp` arguments
leodemoura Dec 29, 2023
dc1d2c8
fix: `simp.trace` missing pre annotation
leodemoura Dec 29, 2023
f7f73de
chore: update stage0
leodemoura Dec 29, 2023
88a0bbc
chore: fix test
leodemoura Dec 29, 2023
f275cce
feat: add option `simprocs`
leodemoura Dec 30, 2023
ea231ef
feat: trace `simproc`s
leodemoura Dec 30, 2023
33bc436
feat: add `simp` option `- <simproc-name>`
leodemoura Dec 30, 2023
04ce796
chore: missing copyright
leodemoura Dec 30, 2023
1f103b7
feat: add builtin `simproc` support
leodemoura Dec 30, 2023
d86771b
chore: update stage0
leodemoura Dec 30, 2023
dd7e6e3
feat: add basic simprocs for Nat
leodemoura Dec 30, 2023
a1d438c
chore: remove bogus `registerSimproc`
leodemoura Dec 31, 2023
4e6ec09
chore: update stage0
leodemoura Dec 31, 2023
e2c49b4
feat: add `simproc`s for `Fin`
leodemoura Dec 31, 2023
c2cbef1
feat: replace `ite` and `dite` shortcircuit theorems with simproc
leodemoura Dec 31, 2023
549d47e
feat: add `simproc`s for `UInt`
leodemoura Jan 1, 2024
df2ecb5
feat: add `simproc`s for `Int`
leodemoura Jan 2, 2024
c8b2d6f
chore: typo
leodemoura Jan 2, 2024
822158d
test: `Int` `simproc`s
leodemoura Jan 2, 2024
bb6a7fa
fix: allow builtin `simproc`s to be provided to `simp` even if they a…
leodemoura Jan 2, 2024
dfe3983
chore: update stage0
leodemoura Jan 2, 2024
382db25
test: builtin `simproc` option that is not in the environment
leodemoura Jan 2, 2024
80b3e87
chore: remove staging workaround
leodemoura Jan 2, 2024
9826229
chore: update stage0
leodemoura Jan 2, 2024
d1c385d
chore: remove staging workaround
leodemoura Jan 2, 2024
33295bc
chore: update stage0
leodemoura Jan 2, 2024
b57fdc0
chore: fix tests
leodemoura Jan 2, 2024
789979e
fix: trace used builtin `simproc`s even if they are not in the enviro…
leodemoura Jan 2, 2024
9c16fed
chore: add another `simproc` test
leodemoura Jan 2, 2024
2d738dc
chore: add default parameter value for `(simprocs : Simprocs)`
leodemoura Jan 2, 2024
8fbf408
chore: better method names
leodemoura Jan 2, 2024
a23029a
chore: use mathlib naming convention
leodemoura Jan 3, 2024
2bab43e
chore: update stage0
leodemoura Jan 3, 2024
f4a213f
chroe: fix tests
leodemoura Jan 3, 2024
7c5a37b
chore: cleanup builtin `simproc`s using `OptionT`
leodemoura Jan 7, 2024
72e5a94
test: test for panic in simprocs
kim-em Jan 7, 2024
e6ed385
fix: panic at `ite` and `dite` simprocs
leodemoura Jan 7, 2024
9b2b4aa
test: timeout in Mathlib.Computability.PartrecCode
kim-em Jan 8, 2024
f296779
doc: add `simproc` release notes
leodemoura Jan 9, 2024
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
69 changes: 62 additions & 7 deletions RELEASES.md
Original file line number Diff line number Diff line change
Expand Up @@ -11,6 +11,61 @@ of each version.
v4.6.0 (development in progress)
---------

* Add custom simplification procedures (aka `simproc`s) to `simp`. Simprocs can be triggered by the simplifier on a specified term-pattern. Here is an small example:
```lean
import Lean.Meta.Tactic.Simp.BuiltinSimprocs.Nat

def foo (x : Nat) : Nat :=
x + 10

/--
The `simproc` `reduceFoo` is invoked on terms that match the pattern `foo _`.
-/
simproc reduceFoo (foo _) :=
/- A term of type `Expr → SimpM (Option Step) -/
fun e => OptionT.run do
/- `simp` uses matching modulo reducibility. So, we ensure the term is a `foo`-application. -/
guard (e.isAppOfArity ``foo 1)
/- `Nat.fromExpr?` tries to convert an expression into a `Nat` value -/
let n ← Nat.fromExpr? e.appArg!
/-
The `Step` type has two constructors: `.done` and `.visit`.
* The constructor `.done` instructs `simp` that the result does
not need to be simplied further.
* The constructor `.visit` instructs `simp` to visit the resulting expression.

If the result holds definitionally as in this example, the field `proof?` can be omitted.
-/
return .done { expr := Lean.mkNatLit (n+10) }
```
We disable simprocs support by using the command `set_option simprocs false`. This command is particularly useful when porting files to v4.6.0.
Simprocs can be scoped, manually added to `simp` commands, and suppressed using `-`. They are also supported by `simp?`. `simp only` does not execute any `simproc`. Here are some examples for the `simproc` defined above.
```lean
example : x + foo 2 = 12 + x := by
set_option simprocs false in
/- This `simp` command does make progress since `simproc`s are disabled. -/
fail_if_success simp
simp_arith

example : x + foo 2 = 12 + x := by
/- `simp only` must not use the default simproc set. -/
fail_if_success simp only
simp_arith

example : x + foo 2 = 12 + x := by
/-
`simp only` does not use the default simproc set,
but we can provide simprocs as arguments. -/
simp only [reduceFoo]
simp_arith

example : x + foo 2 = 12 + x := by
/- We can use `-` to disable `simproc`s. -/
fail_if_success simp [-reduceFoo]
simp_arith
```


v4.5.0
---------

Expand Down Expand Up @@ -50,7 +105,7 @@ v4.5.0
- `Widget.UserWidgetDefinition` is deprecated in favour of `Widget.Module`. The annotation `@[widget]` is deprecated in favour of `@[widget_module]`. To migrate a definition of type `UserWidgetDefinition`, remove the `name` field and replace the type with `Widget.Module`. Removing the `name` results in a title bar no longer being drawn above your panel widget. To add it back, draw it as part of the component using `<details open=true><summary class='mv2 pointer'>{name}</summary>{rest_of_widget}</details>`. See an example migration [here](https://github.com/leanprover/std4/pull/475/files#diff-857376079661a0c28a53b7ff84701afabbdf529836a6944d106c5294f0e68109R43-R83).
- The new command `show_panel_widgets` allows displaying always-on and locally-on panel widgets.
- `RpcEncodable` widget props can now be stored in the infotree.
- See [RFC 2963](https://github.com/leanprover/lean4/issues/2963) for more details and motivation.
- See [RFC 2963](https://github.com/leanprover/lean4/issues/2963) for more details and motivation.

* If no usable lexicographic order can be found automatically for a termination proof, explain why.
See [feat: GuessLex: if no measure is found, explain why](https://github.com/leanprover/lean4/pull/2960).
Expand All @@ -71,16 +126,16 @@ v4.5.0
* Tactics with `withLocation *` [no longer fail](https://github.com/leanprover/lean4/pull/2917) if they close the main goal.

* Implementation of a `test_extern` command for writing tests for `@[extern]` and `@[implemented_by]` functions.
Usage is
Usage is
```
import Lean.Util.TestExtern

test_extern Nat.add 17 37
```
The head symbol must be the constant with the `@[extern]` or `@[implemented_by]` attribute. The return type must have a `DecidableEq` instance.

Bug fixes for
[#2853](https://github.com/leanprover/lean4/issues/2853), [#2953](https://github.com/leanprover/lean4/issues/2953), [#2966](https://github.com/leanprover/lean4/issues/2966),
Bug fixes for
[#2853](https://github.com/leanprover/lean4/issues/2853), [#2953](https://github.com/leanprover/lean4/issues/2953), [#2966](https://github.com/leanprover/lean4/issues/2966),
[#2971](https://github.com/leanprover/lean4/issues/2971), [#2990](https://github.com/leanprover/lean4/issues/2990), [#3094](https://github.com/leanprover/lean4/issues/3094).

Bug fix for [eager evaluation of default value](https://github.com/leanprover/lean4/pull/3043) in `Option.getD`.
Expand All @@ -93,19 +148,19 @@ v4.4.0
---------

* Lake and the language server now support per-package server options using the `moreServerOptions` config field, as well as options that apply to both the language server and `lean` using the `leanOptions` config field. Setting either of these fields instead of `moreServerArgs` ensures that viewing files from a dependency uses the options for that dependency. Additionally, `moreServerArgs` is being deprecated in favor of the `moreGlobalServerArgs` field. See PR [#2858](https://github.com/leanprover/lean4/pull/2858).

A Lakefile with the following deprecated package declaration:
```lean
def moreServerArgs := #[
"-Dpp.unicode.fun=true"
]
def moreLeanArgs := moreServerArgs

package SomePackage where
moreServerArgs := moreServerArgs
moreLeanArgs := moreLeanArgs
```

... can be updated to the following package declaration to use per-package options:
```lean
package SomePackage where
Expand Down
1 change: 1 addition & 0 deletions src/Init.lean
Original file line number Diff line number Diff line change
Expand Up @@ -23,4 +23,5 @@ import Init.NotationExtra
import Init.SimpLemmas
import Init.Hints
import Init.Conv
import Init.Simproc
import Init.SizeOfLemmas
2 changes: 1 addition & 1 deletion src/Init/Data/Fin/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -100,7 +100,7 @@ instance : ShiftLeft (Fin n) where
instance : ShiftRight (Fin n) where
shiftRight := Fin.shiftRight

instance : OfNat (Fin (no_index (n+1))) i where
instance instOfNat : OfNat (Fin (no_index (n+1))) i where
ofNat := Fin.ofNat i

instance : Inhabited (Fin (no_index (n+1))) where
Expand Down
2 changes: 1 addition & 1 deletion src/Init/Data/Int/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -49,7 +49,7 @@ attribute [extern "lean_int_neg_succ_of_nat"] Int.negSucc

instance : Coe Nat Int := ⟨Int.ofNat⟩

instance : OfNat Int n where
instance instOfNat : OfNat Int n where
ofNat := Int.ofNat n

namespace Int
Expand Down
3 changes: 2 additions & 1 deletion src/Init/Data/List/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -124,7 +124,8 @@ def appendTR (as bs : List α) : List α :=
induction as with
| nil => rfl
| cons a as ih =>
simp [reverseAux, List.append, ih, reverseAux_reverseAux]
rw [reverseAux, reverseAux_reverseAux]
simp [List.append, ih, reverseAux]

instance : Append (List α) := ⟨List.append⟩

Expand Down
11 changes: 5 additions & 6 deletions src/Init/Data/UInt/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -39,7 +39,7 @@ def UInt8.shiftRight (a b : UInt8) : UInt8 := ⟨a.val >>> (modn b 8).val⟩
def UInt8.lt (a b : UInt8) : Prop := a.val < b.val
def UInt8.le (a b : UInt8) : Prop := a.val ≤ b.val

instance : OfNat UInt8 n := ⟨UInt8.ofNat n⟩
instance UInt8.instOfNat : OfNat UInt8 n := ⟨UInt8.ofNat n⟩
instance : Add UInt8 := ⟨UInt8.add⟩
instance : Sub UInt8 := ⟨UInt8.sub⟩
instance : Mul UInt8 := ⟨UInt8.mul⟩
Expand Down Expand Up @@ -110,8 +110,7 @@ def UInt16.shiftRight (a b : UInt16) : UInt16 := ⟨a.val >>> (modn b 16).val⟩
def UInt16.lt (a b : UInt16) : Prop := a.val < b.val
def UInt16.le (a b : UInt16) : Prop := a.val ≤ b.val


instance : OfNat UInt16 n := ⟨UInt16.ofNat n⟩
instance UInt16.instOfNat : OfNat UInt16 n := ⟨UInt16.ofNat n⟩
instance : Add UInt16 := ⟨UInt16.add⟩
instance : Sub UInt16 := ⟨UInt16.sub⟩
instance : Mul UInt16 := ⟨UInt16.mul⟩
Expand Down Expand Up @@ -184,7 +183,7 @@ def UInt8.toUInt32 (a : UInt8) : UInt32 := a.toNat.toUInt32
@[extern "lean_uint16_to_uint32"]
def UInt16.toUInt32 (a : UInt16) : UInt32 := a.toNat.toUInt32

instance : OfNat UInt32 n := ⟨UInt32.ofNat n⟩
instance UInt32.instOfNat : OfNat UInt32 n := ⟨UInt32.ofNat n⟩
instance : Add UInt32 := ⟨UInt32.add⟩
instance : Sub UInt32 := ⟨UInt32.sub⟩
instance : Mul UInt32 := ⟨UInt32.mul⟩
Expand Down Expand Up @@ -244,7 +243,7 @@ def UInt16.toUInt64 (a : UInt16) : UInt64 := a.toNat.toUInt64
@[extern "lean_uint32_to_uint64"]
def UInt32.toUInt64 (a : UInt32) : UInt64 := a.toNat.toUInt64

instance : OfNat UInt64 n := ⟨UInt64.ofNat n⟩
instance UInt64.instOfNat : OfNat UInt64 n := ⟨UInt64.ofNat n⟩
instance : Add UInt64 := ⟨UInt64.add⟩
instance : Sub UInt64 := ⟨UInt64.sub⟩
instance : Mul UInt64 := ⟨UInt64.mul⟩
Expand Down Expand Up @@ -322,7 +321,7 @@ def USize.toUInt32 (a : USize) : UInt32 := a.toNat.toUInt32
def USize.lt (a b : USize) : Prop := a.val < b.val
def USize.le (a b : USize) : Prop := a.val ≤ b.val

instance : OfNat USize n := ⟨USize.ofNat n⟩
instance USize.instOfNat : OfNat USize n := ⟨USize.ofNat n⟩
instance : Add USize := ⟨USize.add⟩
instance : Sub USize := ⟨USize.sub⟩
instance : Mul USize := ⟨USize.mul⟩
Expand Down
8 changes: 8 additions & 0 deletions src/Init/SimpLemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -10,6 +10,7 @@ import Init.Core
set_option linter.missingDocs true -- keep it documented

theorem of_eq_true (h : p = True) : p := h ▸ trivial
theorem of_eq_false (h : p = False) : ¬ p := fun hp => False.elim (h.mp hp)

theorem eq_true (h : p) : p = True :=
propext ⟨fun _ => trivial, fun _ => h⟩
Expand Down Expand Up @@ -84,6 +85,13 @@ theorem dite_congr {_ : Decidable b} [Decidable c]
@[simp] theorem ite_false (a b : α) : (if False then a else b) = b := rfl
@[simp] theorem dite_true {α : Sort u} {t : True → α} {e : ¬ True → α} : (dite True t e) = t True.intro := rfl
@[simp] theorem dite_false {α : Sort u} {t : False → α} {e : ¬ False → α} : (dite False t e) = e not_false := rfl
section SimprocHelperLemmas
set_option simprocs false
theorem ite_cond_eq_true {α : Sort u} {c : Prop} {_ : Decidable c} (a b : α) (h : c = True) : (if c then a else b) = a := by simp [h]
theorem ite_cond_eq_false {α : Sort u} {c : Prop} {_ : Decidable c} (a b : α) (h : c = False) : (if c then a else b) = b := by simp [h]
theorem dite_cond_eq_true {α : Sort u} {c : Prop} {_ : Decidable c} {t : c → α} {e : ¬ c → α} (h : c = True) : (dite c t e) = t (of_eq_true h) := by simp [h]
theorem dite_cond_eq_false {α : Sort u} {c : Prop} {_ : Decidable c} {t : c → α} {e : ¬ c → α} (h : c = False) : (dite c t e) = e (of_eq_false h) := by simp [h]
end SimprocHelperLemmas
@[simp] theorem ite_self {α : Sort u} {c : Prop} {d : Decidable c} (a : α) : ite c a a = a := by cases d <;> rfl
@[simp] theorem and_self (p : Prop) : (p ∧ p) = p := propext ⟨(·.1), fun h => ⟨h, h⟩⟩
@[simp] theorem and_true (p : Prop) : (p ∧ True) = p := propext ⟨(·.1), (⟨·, trivial⟩)⟩
Expand Down
94 changes: 94 additions & 0 deletions src/Init/Simproc.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,94 @@
/-
Copyright (c) 2023 Amazon.com, Inc. or its affiliates. All Rights Reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Leonardo de Moura
-/
prelude
import Init.NotationExtra

namespace Lean.Parser
/--
A user-defined simplification procedure used by the `simp` tactic, and its variants.
Here is an example.
```lean
simproc reduce_add (_ + _) := fun e => do
unless (e.isAppOfArity ``HAdd.hAdd 6) do return none
let some n ← getNatValue? (e.getArg! 4) | return none
let some m ← getNatValue? (e.getArg! 5) | return none
return some (.done { expr := mkNatLit (n+m) })
```
The `simp` tactic invokes `reduce_add` whenever it finds a term of the form `_ + _`.
The simplification procedures are stored in an (imperfect) discrimination tree.
The procedure should **not** assume the term `e` perfectly matches the given pattern.
The body of a simplification procedure must have type `Simproc`, which is an alias for
`Expr → SimpM (Option Step)`.
You can instruct the simplifier to apply the procedure before its sub-expressions
have been simplified by using the modifier `↓` before the procedure name. Example.
```lean
simproc ↓ reduce_add (_ + _) := fun e => ...
```
Simplification procedures can be also scoped or local.
-/
syntax (docComment)? attrKind "simproc " (Tactic.simpPre <|> Tactic.simpPost)? ident " (" term ")" " := " term : command

/--
A user-defined simplification procedure declaration. To activate this procedure in `simp` tactic,
we must provide it as an argument, or use the command `attribute` to set its `[simproc]` attribute.
-/
syntax (docComment)? "simproc_decl " ident " (" term ")" " := " term : command

/--
A builtin simplification procedure.
-/
syntax (docComment)? attrKind "builtin_simproc " (Tactic.simpPre <|> Tactic.simpPost)? ident " (" term ")" " := " term : command

/--
A builtin simplification procedure declaration.
-/
syntax (docComment)? "builtin_simproc_decl " ident " (" term ")" " := " term : command

/--
Auxiliary command for associating a pattern with a simplification procedure.
-/
syntax (name := simprocPattern) "simproc_pattern% " term " => " ident : command

/--
Auxiliary command for associating a pattern with a builtin simplification procedure.
-/
syntax (name := simprocPatternBuiltin) "builtin_simproc_pattern% " term " => " ident : command

namespace Attr
/--
Auxiliary attribute for simplification procedures.
-/
syntax (name := simprocAttr) "simproc" (Tactic.simpPre <|> Tactic.simpPost)? : attr

/--
Auxiliary attribute for builtin simplification procedures.
-/
syntax (name := simprocBuiltinAttr) "builtin_simproc" (Tactic.simpPre <|> Tactic.simpPost)? : attr
end Attr

macro_rules
| `($[$doc?:docComment]? simproc_decl $n:ident ($pattern:term) := $body) => do
let simprocType := `Lean.Meta.Simp.Simproc
`($[$doc?:docComment]? def $n:ident : $(mkIdent simprocType) := $body
simproc_pattern% $pattern => $n)

macro_rules
| `($[$doc?:docComment]? builtin_simproc_decl $n:ident ($pattern:term) := $body) => do
let simprocType := `Lean.Meta.Simp.Simproc
`($[$doc?:docComment]? def $n:ident : $(mkIdent simprocType) := $body
builtin_simproc_pattern% $pattern => $n)

macro_rules
| `($[$doc?:docComment]? $kind:attrKind simproc $[$pre?]? $n:ident ($pattern:term) := $body) => do
`(simproc_decl $n ($pattern) := $body
attribute [$kind simproc $[$pre?]?] $n)

macro_rules
| `($[$doc?:docComment]? $kind:attrKind builtin_simproc $[$pre?]? $n:ident ($pattern:term) := $body) => do
`(builtin_simproc_decl $n ($pattern) := $body
attribute [$kind builtin_simproc $[$pre?]?] $n)

end Lean.Parser
2 changes: 1 addition & 1 deletion src/Init/Tactics.lean
Original file line number Diff line number Diff line change
Expand Up @@ -753,7 +753,7 @@ end Tactic

namespace Attr
/--
Theorems tagged with the `simp` attribute are by the simplifier
Theorems tagged with the `simp` attribute are used by the simplifier
(i.e., the `simp` tactic, and its variants) to simplify expressions occurring in your goals.
We call theorems tagged with the `simp` attribute "simp theorems" or "simp lemmas".
Lean maintains a database/index containing all active simp theorems.
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Elab/PreDefinition/Structural/Eqns.lean
Original file line number Diff line number Diff line change
Expand Up @@ -42,7 +42,7 @@ where
go mvarId
else if let some mvarId ← whnfReducibleLHS? mvarId then
go mvarId
else match (← simpTargetStar mvarId {}).1 with
else match (← simpTargetStar mvarId {} (simprocs := {})).1 with
| TacticResultCNM.closed => return ()
| TacticResultCNM.modified mvarId => go mvarId
| TacticResultCNM.noChange =>
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Elab/PreDefinition/WF/Eqns.lean
Original file line number Diff line number Diff line change
Expand Up @@ -122,7 +122,7 @@ where
match (← reduceRecMatcher? e) with
| some e' => return Simp.Step.done { expr := e' }
| none =>
match (← Simp.simpMatchCore? app e SplitIf.discharge?) with
match (← Simp.simpMatchCore? app.matcherName e SplitIf.discharge?) with
| some r => return r
| none => return Simp.Step.visit { expr := e }

Expand Down Expand Up @@ -169,7 +169,7 @@ private partial def mkProof (declName : Name) (info : EqnInfo) (type : Expr) : M
go mvarId
else if let some mvarId ← whnfReducibleLHS? mvarId then
go mvarId
else match (← simpTargetStar mvarId { config.dsimp := false }).1 with
else match (← simpTargetStar mvarId { config.dsimp := false } (simprocs := {})).1 with
| TacticResultCNM.closed => return ()
| TacticResultCNM.modified mvarId => go mvarId
| TacticResultCNM.noChange =>
Expand Down
1 change: 1 addition & 0 deletions src/Lean/Elab/Tactic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -13,6 +13,7 @@ import Lean.Elab.Tactic.Match
import Lean.Elab.Tactic.Rewrite
import Lean.Elab.Tactic.Location
import Lean.Elab.Tactic.Simp
import Lean.Elab.Tactic.Simproc
import Lean.Elab.Tactic.BuiltinTactic
import Lean.Elab.Tactic.Split
import Lean.Elab.Tactic.Conv
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Elab/Tactic/Conv/Simp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -17,9 +17,9 @@ def applySimpResult (result : Simp.Result) : TacticM Unit := do
updateLhs result.expr (← result.getProof)

@[builtin_tactic Lean.Parser.Tactic.Conv.simp] def evalSimp : Tactic := fun stx => withMainContext do
let { ctx, dischargeWrapper, .. } ← mkSimpContext stx (eraseLocal := false)
let { ctx, simprocs, dischargeWrapper, .. } ← mkSimpContext stx (eraseLocal := false)
let lhs ← getLhs
let (result, _) ← dischargeWrapper.with fun d? => simp lhs ctx (discharge? := d?)
let (result, _) ← dischargeWrapper.with fun d? => simp lhs ctx (simprocs := simprocs) (discharge? := d?)
applySimpResult result

@[builtin_tactic Lean.Parser.Tactic.Conv.simpMatch] def evalSimpMatch : Tactic := fun _ => withMainContext do
Expand Down
Loading
Loading