Skip to content

Commit 823671f

Browse files
authored
feat: set_option tactic in grind interactive mode (#10843)
This PR implements the `set_option` tactic in `grind` interactive mode.
1 parent 681724a commit 823671f

File tree

3 files changed

+29
-0
lines changed

3 files changed

+29
-0
lines changed

src/Init/Grind/Interactive.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -189,5 +189,10 @@ generated `grind` tactic scripts.
189189
-/
190190
syntax (name := exposeNames) "expose_names" : grind
191191

192+
/--
193+
`set_option opt val in tacs` (the tactic) acts like `set_option opt val` at the command level,
194+
but it sets the option only within the tactics `tacs`. -/
195+
syntax (name := setOption) "set_option " ident ppSpace optionValue " in " grindSeq : grind
196+
192197
end Grind
193198
end Lean.Parser.Tactic

src/Lean/Elab/Tactic/Grind/BuiltinTactic.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -25,6 +25,7 @@ import Lean.Elab.Tactic.Basic
2525
import Lean.Elab.Tactic.RenameInaccessibles
2626
import Lean.Elab.Tactic.Grind.Filter
2727
import Lean.Elab.Tactic.Grind.ShowState
28+
import Lean.Elab.SetOption
2829
namespace Lean.Elab.Tactic.Grind
2930

3031
def showStateAt (ref : Syntax) (filter : Filter) : GrindTacticM Unit := do
@@ -424,4 +425,8 @@ where
424425
liftGrindM <| resetAnchors
425426
replaceMainGoal [{ goal with mvarId }]
426427

428+
@[builtin_grind_tactic setOption] def elabSetOption : GrindTactic := fun stx => do
429+
let options ← Elab.elabSetOption stx[1] stx[2]
430+
withOptions (fun _ => options) do evalGrindTactic stx[4]
431+
427432
end Lean.Elab.Tactic.Grind

tests/lean/run/grind_interactive.lean

Lines changed: 19 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -431,6 +431,25 @@ example (as bs cs : Array α) (v₁ v₂ : α)
431431
grind =>
432432
repeat instantiate only [= Array.getElem_set]
433433

434+
/--
435+
trace: [grind.ematch.instance] Array.getElem_set: (as.set i₁ v₁ ⋯)[j] = if i₁ = j then v₁ else as[j]
436+
-/
437+
#guard_msgs in
438+
example (as bs cs : Array α) (v₁ v₂ : α)
439+
(i₁ i₂ j : Nat)
440+
(h₁ : i₁ < as.size)
441+
(h₂ : bs = as.set i₁ v₁)
442+
(h₃ : i₂ < bs.size)
443+
(h₃ : cs = bs.set i₂ v₂)
444+
(h₄ : i₁ ≠ j ∧ i₂ ≠ j)
445+
(h₅ : j < cs.size)
446+
(h₆ : j < as.size)
447+
: cs[j] = as[j] := by
448+
grind =>
449+
instantiate
450+
set_option trace.grind.ematch.instance true in
451+
instantiate
452+
434453
opaque p : Nat → Prop
435454
opaque q : Nat → Prop
436455
opaque f : Nat → Nat

0 commit comments

Comments
 (0)