Skip to content

Commit 78b09d5

Browse files
authored
feat: support case label like syntax in mvcgen invariants (#10570)
This PR adds support for case label like syntax in `mvcgen invariants` in order to refer to inaccessible names. Example: ```lean def copy (l : List Nat) : Id (Array Nat) := do let mut acc := #[] for x in l do acc := acc.push x return acc theorem copy_labelled_invariants (l : List Nat) : ⦃⌜True⌝⦄ copy l ⦃⇓ r => ⌜r = l.toArray⌝⦄ := by mvcgen [copy] invariants | inv1 acc => ⇓ ⟨xs, letMuts⟩ => ⌜acc = l.toArray⌝ with admit ```
1 parent a164ae5 commit 78b09d5

File tree

4 files changed

+85
-17
lines changed

4 files changed

+85
-17
lines changed

src/Init/Tactics.lean

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2217,6 +2217,15 @@ mvcgen [...] invariants
22172217
· I2
22182218
with grind
22192219
```
2220+
When `I1` and `I2` need to refer to inaccessibles (`mvcgen` will introduce a lot of them for program
2221+
variables), you can use case label syntax:
2222+
```
2223+
mvcgen [...] invariants
2224+
| inv1 _ acc _ => I1 acc
2225+
| _ => I2
2226+
with grind
2227+
```
2228+
This is more convenient than the equivalent `· by rename_i _ acc _; exact I1 acc`.
22202229
22212230
### Invariant suggestions
22222231

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

Lines changed: 34 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -370,15 +370,40 @@ def elabInvariants (stx : Syntax) (invariants : Array MVarId) (suggestInvariant
370370
match stx with
371371
| `(invariantAlts| $invariantsKW $alts*) =>
372372
let invariants ← invariants.filterM (not <$> ·.isAssigned)
373+
374+
let mut dotOrCase := LBool.undef -- .true => dot
373375
for h : n in 0...alts.size do
374-
let alt := alts[n]
375-
match alt with
376-
| `(invariantAlt| · $rhs) =>
377-
let some mv := invariants[n]? | do
378-
logErrorAt rhs m!"More invariants have been defined ({alts.size}) than there were unassigned invariants goals `inv<n>` ({invariants.size})."
379-
continue
380-
discard <| evalTacticAt (← `(tactic| exact $rhs)) mv
381-
| _ => logErrorAt alt m!"Expected invariantAlt, got {alt}"
376+
let alt := alts[n]
377+
match alt with
378+
| `(invariantDotAlt| · $rhs) =>
379+
if dotOrCase matches .false then
380+
logErrorAt alt m!"Alternation between labelled and bulleted invariants is not supported."
381+
break
382+
dotOrCase := .true
383+
let some mv := invariants[n]? | do
384+
logErrorAt alt m!"More invariants have been defined ({alts.size}) than there were unassigned invariants goals `inv<n>` ({invariants.size})."
385+
continue
386+
withRef rhs do
387+
discard <| evalTacticAt (← `(tactic| exact $rhs)) mv
388+
| `(invariantCaseAlt| | $tag $args* => $rhs) =>
389+
if dotOrCase matches .true then
390+
logErrorAt alt m!"Alternation between labelled and bulleted invariants is not supported."
391+
break
392+
dotOrCase := .false
393+
let n? : Option Nat := do
394+
let `(binderIdent| $tag:ident) := tag | some n -- fall back to ordinal
395+
let .str .anonymous s := tag.getId | none
396+
s.dropPrefix? "inv" >>= Substring.toNat?
397+
let some mv := do invariants[(← n?) - 1]? | do
398+
logErrorAt alt m!"No invariant with label {tag} {repr tag}."
399+
continue
400+
if ← mv.isAssigned then
401+
logErrorAt alt m!"Invariant {n?.get!} is already assigned."
402+
continue
403+
withRef rhs do
404+
discard <| evalTacticAt (← `(tactic| rename_i $args*; exact $rhs)) mv
405+
| _ => logErrorAt alt m!"Expected `invariantDotAlt`, got {alt}"
406+
382407
if let `(invariantsKW| invariants) := invariantsKW then
383408
if alts.size < invariants.size then
384409
let missingTypes ← invariants[alts.size:].toArray.mapM (·.getType)
@@ -391,7 +416,7 @@ def elabInvariants (stx : Syntax) (invariants : Array MVarId) (suggestInvariant
391416
if ← mv.isAssigned then
392417
continue
393418
let invariant ← suggestInvariant mv
394-
suggestions := suggestions.push (← `(invariantAlt| · $invariant))
419+
suggestions := suggestions.push (← `(invariantDotAlt| · $invariant))
395420
let alts' := alts ++ suggestions
396421
let stx' ← `(invariantAlts|invariants $alts'*)
397422
if suggestions.size > 0 then

src/Std/Tactic/Do/Syntax.lean

Lines changed: 13 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -328,7 +328,12 @@ macro "mvcgen_trivial" : tactic =>
328328
/--
329329
An invariant alternative of the form `· term`, one per invariant goal.
330330
-/
331-
syntax invariantAlt := ppDedent(ppLine) cdotTk (colGe term)
331+
syntax invariantDotAlt := ppDedent(ppLine) cdotTk (colGe term)
332+
333+
/--
334+
An invariant alternative of the form `| inv<n> a b c => term`, one per invariant goal.
335+
-/
336+
syntax invariantCaseAlt := ppDedent(ppLine) "| " caseArg " => " (colGe term)
332337

333338
/--
334339
Either the contextual keyword ` invariants ` or its tracing form ` invariants? ` which suggests
@@ -337,11 +342,14 @@ skeletons for missing invariants as a hint.
337342
syntax invariantsKW := &"invariants " <|> &"invariants? "
338343

339344
/--
340-
After `mvcgen [...]`, there can be an optional `invariants` followed by a bulleted list of
341-
invariants `· term; · term`.
342-
The tracing variant ` invariants? ` will suggest a skeleton for missing invariants.
345+
After `mvcgen [...]`, there can be an optional `invariants` followed by either
346+
* a bulleted list of invariants `· term; · term`.
347+
* a labelled list of invariants `| inv1 => term; inv2 a b c => term`, which is useful for naming
348+
inaccessibles.
349+
The tracing variant ` invariants? ` will suggest a skeleton for missing invariants; see the
350+
docstring for `mvcgen`.
343351
-/
344-
syntax invariantAlts := invariantsKW withPosition((colGe invariantAlt)*)
352+
syntax invariantAlts := invariantsKW withPosition((colGe (invariantDotAlt <|> invariantCaseAlt))*)
345353

346354
/--
347355
In induction alternative, which can have 1 or more cases on the left

tests/lean/run/mvcgenInvariantsWith.lean

Lines changed: 29 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -94,15 +94,15 @@ theorem test_with_pretac {m : Option Nat} (h : m = some 4) :
9494
(match m with
9595
| some n => (set n : StateM Nat PUnit)
9696
| none => set 0)
97-
⦃⇓ r s => ⌜s = 4⌝⦄ := by
97+
⦃⇓ _ s => ⌜s = 4⌝⦄ := by
9898
mvcgen with simp_all
9999

100100
theorem test_with_cases {m : Option Nat} (h : m = some 4) :
101101
⦃⌜True⌝⦄
102102
(match m with
103103
| some n => (set n : StateM Nat PUnit)
104104
| none => set 0)
105-
⦃⇓ r s => ⌜s = 4⌝⦄ := by
105+
⦃⇓ _ s => ⌜s = 4⌝⦄ := by
106106
mvcgen
107107
with
108108
| vc1 => grind
@@ -113,7 +113,7 @@ theorem test_with_pretac_cases {m : Option Nat} (h : m = some 4) :
113113
(match m with
114114
| some n => (set n : StateM Nat PUnit)
115115
| none => set 0)
116-
⦃⇓ r s => ⌜s = 4⌝⦄ := by
116+
⦃⇓ _ s => ⌜s = 4⌝⦄ := by
117117
mvcgen
118118
with simp -- `simp` is a no-op on some goals, but it should not fail
119119
| vc1 => grind
@@ -193,3 +193,29 @@ theorem nodup_twice_missing_two_invariants (l : List Int) : nodup_twice l ↔ l.
193193
mvcgen
194194
invariants
195195
with grind
196+
197+
def copy (l : List Nat) : Id (Array Nat) := do
198+
let mut acc := #[]
199+
for x in l do
200+
acc := acc.push x
201+
return acc
202+
203+
set_option warn.sorry false in
204+
theorem copy_labelled_invariants (l : List Nat) : ⦃⌜True⌝⦄ copy l ⦃⇓ r => ⌜r = l.toArray⌝⦄ := by
205+
mvcgen [copy] invariants
206+
| inv1 acc => ⇓ ⟨xs, letMuts⟩ => ⌜acc = l.toArray⌝
207+
with admit
208+
209+
set_option warn.sorry false in
210+
theorem copy_labelled_invariants_noname (l : List Nat) : ⦃⌜True⌝⦄ copy l ⦃⇓ r => ⌜r = l.toArray⌝⦄ := by
211+
mvcgen [copy] invariants
212+
| _ acc => ⇓ ⟨xs, letMuts⟩ => ⌜acc = l.toArray⌝
213+
with admit
214+
215+
/-- error: Alternation between labelled and bulleted invariants is not supported. -/
216+
#guard_msgs in
217+
theorem copy_labelled_invariants_dontmix (l : List Nat) : ⦃⌜True⌝⦄ copy l ⦃⇓ r => ⌜r = l.toArray⌝⦄ := by
218+
mvcgen [copy] invariants
219+
· ⇓ ⟨xs, letMuts⟩ => ⌜True⌝
220+
| _ acc => ⇓ ⟨xs, letMuts⟩ => ⌜acc = l.toArray⌝
221+
with admit

0 commit comments

Comments
 (0)