Skip to content

Commit 9859494

Browse files
authored
fix: try synthesizing synthetic MVars in mspec (#10644)
This PR explicitly tries to synthesize synthetic MVars in `mspec`. Doing so resolves a bug triggered by use of the loop invariant lemma for `Std.PRange`.
1 parent ba52e93 commit 9859494

File tree

4 files changed

+96
-31
lines changed

4 files changed

+96
-31
lines changed

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

Lines changed: 17 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -14,15 +14,13 @@ public import Lean.Elab.Tactic.Do.ProofMode.Assumption
1414
public import Lean.Elab.Tactic.Do.Attr
1515
public import Std.Do.Triple
1616

17-
public section
18-
1917
namespace Lean.Elab.Tactic.Do
2018
open Lean Elab Tactic Meta
2119
open Std.Do Do.SpecAttr Do.ProofMode
2220

2321
builtin_initialize registerTraceClass `Elab.Tactic.Do.spec
2422

25-
def findSpec (database : SpecTheorems) (wp : Expr) : MetaM SpecTheorem := do
23+
public def findSpec (database : SpecTheorems) (wp : Expr) : MetaM SpecTheorem := do
2624
let_expr [email protected] _m _ps _instWP _α prog := wp | throwError "target not a wp application {wp}"
2725
let prog ← instantiateMVarsIfMVarApp prog
2826
let prog := prog.headBeta
@@ -139,7 +137,7 @@ def dischargeMGoal (goal : MGoal) (goalTag : Name) : n Expr := do
139137
/--
140138
Returns the proof and the list of new unassigned MVars.
141139
-/
142-
def mSpec (goal : MGoal) (elabSpecAtWP : Expr → n SpecTheorem) (goalTag : Name) : n Expr := do
140+
public def mSpec (goal : MGoal) (elabSpecAtWP : Expr → n SpecTheorem) (goalTag : Name) : n Expr := do
143141
-- First instantiate `fun s => ...` in the target via repeated `mintro ∀s`.
144142
mIntroForallN goal goal.target.consumeMData.getNumHeadLambdas fun goal => do
145143
-- Elaborate the spec for the wp⟦e⟧ app in the target
@@ -177,11 +175,25 @@ def mSpec (goal : MGoal) (elabSpecAtWP : Expr → n SpecTheorem) (goalTag : Name
177175
let excessArgs := (args.extract 4 args.size).reverse
178176

179177
-- Actually instantiate the specThm using the expected type computed from `wp`.
180-
let_expr f@Triple m ps instWP α prog P Q := specTy | do liftMetaM (throwError "target not a Triple application {specTy}")
178+
let_expr f@Triple m ps instWP α prog P Q := specTy
179+
| liftMetaM <| throwError "target not a Triple application {specTy}"
181180
let wp' := mkApp5 (mkConst ``WP.wp f.constLevels!) m ps instWP α prog
182181
unless (← withAssignableSyntheticOpaque <| isDefEq wp wp') do
183182
Term.throwTypeMismatchError none wp wp' spec
184183

184+
-- Try synthesizing synthetic MVars. We don't have the convenience of `TermElabM`, hence
185+
-- this poor man's version of `TermElabM.synthesizeSyntheticMVars`.
186+
-- We do so after the def eq call so that instance resolution is likely to succeed.
187+
-- If it _doesn't_ succeed now, then the spec theorem leaves behind an additional subgoal.
188+
-- We'll add a trace message if that happens.
189+
for mvar in mvars do
190+
let mvar := mvar.mvarId!
191+
if (← mvar.getKind) matches .synthetic then
192+
match ← trySynthInstance (← mvar.getType) with
193+
| .some prf => liftMetaM <| mvar.assign prf
194+
| .none => continue
195+
| .undef => liftMetaM <| do trace[Elab.Tactic.Do.spec] "Failed to synthesize synthetic MVar {mvar} from unifying {specTy} against {prog}.\nThis likely leaves behind an additional subgoal."
196+
185197
let P ← instantiateMVarsIfMVarApp P
186198
let Q ← instantiateMVarsIfMVarApp Q
187199

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

Lines changed: 10 additions & 26 deletions
Original file line numberDiff line numberDiff line change
@@ -64,41 +64,25 @@ where
6464
-- trace[Elab.Tactic.Do.vcgen] "fail {goal.toExpr}"
6565
emitVC goal.toExpr name
6666

67-
tryGoal (goal : Expr) (name : Name) : VCGenM Expr := do
68-
-- trace[Elab.Tactic.Do.vcgen] "tryGoal: {goal}"
69-
forallTelescope goal fun xs body => do
70-
let res ← try mStart body catch _ =>
71-
-- trace[Elab.Tactic.Do.vcgen] "not an MGoal: {body}"
72-
return ← mkLambdaFVars xs (← emitVC body name)
67+
tryGoal (mvar : MVarId) : OptionT VCGenM Expr := mvar.withContext do
68+
-- The type might contain more `P ⊢ₛ wp⟦prog⟧ Q` apps. Try and prove it!
69+
forallTelescope (← mvar.getType) fun xs body => do
70+
let res ← try mStart body catch _ => OptionT.fail
7371
-- trace[Elab.Tactic.Do.vcgen] "an MGoal: {res.goal.toExpr}"
74-
let mut prf ← onGoal res.goal name
72+
let mut prf ← onGoal res.goal (← mvar.getTag)
7573
-- res.goal.checkProof prf
7674
if let some proof := res.proof? then
7775
prf := mkApp proof prf
78-
mkLambdaFVars xs prf
76+
returnmkLambdaFVars xs prf
7977

8078
assignMVars (mvars : List MVarId) : VCGenM PUnit := do
8179
for mvar in mvars do
8280
if ← mvar.isAssigned then continue
83-
mvar.withContext <| do
8481
-- trace[Elab.Tactic.Do.vcgen] "assignMVars {← mvar.getTag}, isDelayedAssigned: {← mvar.isDelayedAssigned},\n{mvar}"
85-
let ty ← mvar.getType
86-
if ← isProp ty then
87-
-- Might contain more `P ⊢ₛ wp⟦prog⟧ Q` apps. Try and prove it!
88-
let prf ← tryGoal ty (← mvar.getTag)
89-
if ← mvar.isAssigned then
90-
throwError "Tried to assign already assigned metavariable `{← mvar.getTag}` while `tryGoal`. MVar: {mvar}\nAssignment: {mkMVar mvar}\nNew assignment: {prf}"
91-
mvar.assign prf
92-
return
93-
if ty.isAppOf ``PostCond || ty.isAppOf ``Invariant || ty.isAppOf ``SPred then
94-
-- Here we make `mvar` a synthetic opaque goal upon discharge failure.
95-
-- This is the right call for (previously natural) holes such as loop invariants, which
96-
-- would otherwise lead to spurious instantiations and unwanted renamings (when leaving the
97-
-- scope of a local).
98-
-- But it's wrong for, e.g., schematic variables. The latter should never be PostConds,
99-
-- Invariants or SPreds, hence the condition.
100-
mvar.setKind .syntheticOpaque
101-
addSubGoalAsVC mvar
82+
let some prf ← (tryGoal mvar).run | addSubGoalAsVC mvar
83+
if ← mvar.isAssigned then
84+
throwError "Tried to assign already assigned metavariable `{← mvar.getTag}`. MVar: {mvar}\nAssignment: {mkMVar mvar}\nNew assignment: {prf}"
85+
mvar.assign prf
10286

10387
onGoal goal name : VCGenM Expr := do
10488
let T := goal.target

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

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -94,6 +94,14 @@ def ifOutOfFuel (x : VCGenM α) (k : VCGenM α) : VCGenM α := do
9494
def addSubGoalAsVC (goal : MVarId) : VCGenM PUnit := do
9595
goal.freshenLCtxUserNamesSinceIdx (← read).initialCtxSize
9696
let ty ← goal.getType
97+
if ty.isAppOf ``Std.Do.PostCond || ty.isAppOf ``Std.Do.SPred then
98+
-- Here we make `mvar` a synthetic opaque goal upon discharge failure.
99+
-- This is the right call for (previously natural) holes such as loop invariants, which
100+
-- would otherwise lead to spurious instantiations and unwanted renamings (when leaving the
101+
-- scope of a local).
102+
-- But it's wrong for, e.g., schematic variables. The latter should never be PostConds,
103+
-- Invariants or SPreds, hence the condition.
104+
goal.setKind .syntheticOpaque
97105
if ty.isAppOf ``Std.Do.Invariant then
98106
modify fun s => { s with invariants := s.invariants.push goal }
99107
else
Lines changed: 61 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,61 @@
1+
import Std.Tactic.Do
2+
3+
open Std.Do
4+
5+
set_option warn.sorry false
6+
set_option mvcgen.warning false
7+
8+
def foo (l : List Nat) : Nat := Id.run do
9+
let mut out := 0
10+
for _ in l do
11+
out := out + 1
12+
return out
13+
14+
def bar (n : Nat) : Nat := Id.run do
15+
let mut out := 0
16+
for _ in 0...n do
17+
out := out + 1
18+
return out
19+
20+
/-- This works as expected: -/
21+
theorem foo_eq (l : List Nat) : foo l = l.length := by
22+
generalize h : foo l = x
23+
apply Id.of_wp_run_eq h
24+
mvcgen
25+
case inv1 =>
26+
exact ⇓⟨xs, out⟩ => ⌜True⌝
27+
case vc1 => sorry
28+
case vc2 => sorry
29+
case vc3 => sorry
30+
31+
/-- This used to fail in `inv1` due to an unresolved `HasFiniteRanges` instance: -/
32+
theorem bar_eq (n : Nat) : bar n = n := by
33+
generalize h : bar n = x
34+
apply Id.of_wp_run_eq h
35+
mvcgen
36+
case inv1 =>
37+
-- Invalid match expression: The type of pattern variable 'xs' contains metavariables:
38+
-- (0...n).toList.Cursor
39+
exact ⇓⟨xs, out⟩ => ⌜True⌝
40+
case vc1 => sorry
41+
case vc2 => sorry
42+
case vc3 => sorry
43+
44+
theorem bar_eq' (n : Nat) : bar n = n := by
45+
generalize h : bar n = x
46+
apply Id.of_wp_run_eq h
47+
mvcgen
48+
case inv1 =>
49+
exact ⇓⟨xs, out⟩ => ⌜True⌝
50+
case vc1 => sorry
51+
case vc2 => sorry
52+
case vc3 => sorry
53+
54+
/-- Check the same for `mspec` -/
55+
theorem bar_eq_mspec (n : Nat) : bar n = n := by
56+
generalize h : bar n = x
57+
apply Id.of_wp_run_eq h
58+
mspec
59+
-- should not produce a goal for `Std.PRange.HasFiniteRanges Std.PRange.BoundShape.open Nat`
60+
fail_if_success case inst => exact inferInstance
61+
all_goals sorry

0 commit comments

Comments
 (0)