Skip to content

Commit f3f48d3

Browse files
committed
feat: implement mvcgen?, expanding to mvcgen invariants?
1 parent 746206c commit f3f48d3

File tree

3 files changed

+39
-1
lines changed

3 files changed

+39
-1
lines changed

src/Lean/Elab/Tactic/Do/VCGen.lean

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -478,3 +478,12 @@ def elabMVCGen : Tactic := fun stx => withMainContext do
478478
trace[Elab.Tactic.Do.vcgen] "before replacing main goal {← (invariants ++ vcs).mapM fun m => m.getTag}"
479479
replaceMainGoal (invariants ++ vcs).toList
480480
-- trace[Elab.Tactic.Do.vcgen] "replaced main goal, new: {← getGoals}"
481+
482+
@[builtin_tactic Lean.Parser.Tactic.mvcgenHint]
483+
def elabMVCGenHint : Tactic := fun stx => withMainContext do
484+
let stx' : TSyntax ``mvcgen := TSyntax.mk <| stx
485+
|>.setKind ``Lean.Parser.Tactic.mvcgen
486+
|>.modifyArgs (·.set! 0 (mkAtom "mvcgen") |>.push (mkNullNode #[← `(invariantAlts| invariants?)]) |>.push mkNullNode)
487+
-- logInfo m!"{stx}\n{toString stx}\n{repr stx}"
488+
-- logInfo m!"{stx'}\n{toString stx'}\n{repr stx'}"
489+
Lean.Meta.Tactic.TryThis.addSuggestion stx stx'

src/Std/Tactic/Do/Syntax.lean

Lines changed: 7 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -365,5 +365,11 @@ syntax vcAlts := "with " (ppSpace colGt tactic)? withPosition((colGe vcAlt)*)
365365

366366
@[tactic_alt Lean.Parser.Tactic.mvcgenMacro]
367367
syntax (name := mvcgen) "mvcgen" optConfig
368-
(" [" withoutPosition((simpStar <|> simpErase <|> simpLemma),*,?) "]")?
368+
(" [" withoutPosition((simpStar <|> simpErase <|> simpLemma),*,?) "] ")?
369369
(invariantAlts)? (vcAlts)? : tactic
370+
371+
/--
372+
A hint tactic that expands to `mvcgen invariants?`.
373+
-/
374+
syntax (name := mvcgenHint) "mvcgen?" optConfig
375+
(" [" withoutPosition((simpStar <|> simpErase <|> simpLemma),*,?) "] ")? : tactic

tests/lean/run/mvcgenInvariantsSuggestions.lean

Lines changed: 23 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,29 @@ theorem mySum_suggest_invariant (l : List Nat) : mySum l = l.sum := by
2626
mvcgen invariants?
2727
all_goals admit
2828

29+
/--
30+
info: Try this:
31+
[apply] mvcgen invariants?
32+
---
33+
info: Try this:
34+
[apply] mvcgen [mySum] invariants?
35+
---
36+
info: Try this:
37+
[apply] mvcgen +elimLets invariants?
38+
---
39+
info: Try this:
40+
[apply] mvcgen +elimLets [mySum] invariants?
41+
-/
42+
#guard_msgs (info) in
43+
theorem mySum_suggest_invariant_short (l : List Nat) : mySum l = l.sum := by
44+
generalize h : mySum l = r
45+
apply Id.of_wp_run_eq h
46+
mvcgen?
47+
mvcgen? [mySum]
48+
mvcgen? +elimLets
49+
mvcgen? +elimLets [mySum]
50+
all_goals admit
51+
2952
def mySum2 (l : List Nat) : Nat := Id.run do
3053
let mut acc := 0
3154
let mut acc2 := 0

0 commit comments

Comments
 (0)