Skip to content

Commit 14ff08d

Browse files
authored
feat: repeat tactical for grind interactive mode (#10748)
This PR implements the `repeat` tactical for the `grind` interactive mode.
1 parent 316859e commit 14ff08d

File tree

2 files changed

+24
-0
lines changed

2 files changed

+24
-0
lines changed

src/Init/Grind/Interactive.lean

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -158,5 +158,16 @@ macro "admit" : grind => `(grind| sorry)
158158
/-- `fail msg` is a tactic that always fails, and produces an error using the given message. -/
159159
syntax (name := fail) "fail" (ppSpace str)? : grind
160160

161+
/--
162+
`repeat tac` repeatedly applies `tac` so long as it succeeds.
163+
The tactic `tac` may be a tactic sequence, and if `tac` fails at any point in its execution,
164+
`repeat` will revert any partial changes that `tac` made to the tactic state.
165+
The tactic `tac` should eventually fail, otherwise `repeat tac` will run indefinitely.
166+
-/
167+
syntax "repeat " grindSeq : grind
168+
169+
macro_rules
170+
| `(grind| repeat $seq) => `(grind| first | ($seq); repeat $seq | skip)
171+
161172
end Grind
162173
end Lean.Parser.Tactic

tests/lean/run/grind_interactive.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -376,6 +376,19 @@ example (as bs cs : Array α) (v₁ v₂ : α)
376376
instantiate = Array.getElem_set
377377
instantiate ← Array.getElem_set
378378

379+
example (as bs cs : Array α) (v₁ v₂ : α)
380+
(i₁ i₂ j : Nat)
381+
(h₁ : i₁ < as.size)
382+
(h₂ : bs = as.set i₁ v₁)
383+
(h₃ : i₂ < bs.size)
384+
(h₃ : cs = bs.set i₂ v₂)
385+
(h₄ : i₁ ≠ j ∧ i₂ ≠ j)
386+
(h₅ : j < cs.size)
387+
(h₆ : j < as.size)
388+
: cs[j] = as[j] := by
389+
grind =>
390+
repeat instantiate =Array.getElem_set
391+
379392
opaque p : Nat → Prop
380393
opaque q : Nat → Prop
381394
opaque f : Nat → Nat

0 commit comments

Comments
 (0)