Skip to content

Commit 406bc29

Browse files
committed
fix: hygiene for goals generated by mvcgen
1 parent dfd3d18 commit 406bc29

File tree

2 files changed

+15
-7
lines changed

2 files changed

+15
-7
lines changed

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

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -92,17 +92,17 @@ def ifOutOfFuel (x : VCGenM α) (k : VCGenM α) : VCGenM α := do
9292
| _ => k
9393

9494
def addSubGoalAsVC (goal : MVarId) : VCGenM PUnit := do
95+
goal.freshenLCtxUserNamesSinceIdx (← read).initialCtxSize
9596
let ty ← goal.getType
9697
if ty.isAppOf ``Std.Do.Invariant then
9798
modify fun s => { s with invariants := s.invariants.push goal }
9899
else
99100
modify fun s => { s with vcs := s.vcs.push goal }
100101

101102
def emitVC (subGoal : Expr) (name : Name) : VCGenM Expr := do
102-
withFreshUserNamesSinceIdx (← read).initialCtxSize do
103-
let m ← liftM <| mkFreshExprSyntheticOpaqueMVar subGoal (tag := name)
104-
addSubGoalAsVC m.mvarId!
105-
return m
103+
let m ← liftM <| mkFreshExprSyntheticOpaqueMVar subGoal (tag := name)
104+
addSubGoalAsVC m.mvarId!
105+
return m
106106

107107
def liftSimpM (x : SimpM α) : VCGenM α := do
108108
let ctx ← read

src/Lean/Meta/Basic.lean

Lines changed: 11 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -1968,14 +1968,14 @@ def withErasedFVars [MonadLCtx n] [MonadLiftT MetaM n] (fvarIds : Array FVarId)
19681968
/--
19691969
Ensures that the user names of all local declarations after index `idx` have a macro scope.
19701970
-/
1971-
def withFreshUserNamesSinceIdx [MonadLCtx n] [MonadLiftT MetaM n] (idx : Nat) (k : n α) : n α := do
1972-
let mut lctx ← getLCtx
1971+
private def freshenUserNamesSinceIdx [MonadLiftT MetaM n] (lctx : LocalContext) (idx : Nat) : n LocalContext := do
1972+
let mut lctx := lctx
19731973
for i in [idx:lctx.numIndices] do
19741974
let some decl := lctx.decls[i]! | continue
19751975
let n := decl.userName
19761976
if !n.hasMacroScopes then
19771977
lctx := lctx.setUserName decl.fvarId (← liftMetaM <| mkFreshUserName n)
1978-
withLCtx' lctx k
1978+
return lctx
19791979

19801980
private def withMVarContextImp (mvarId : MVarId) (x : MetaM α) : MetaM α := do
19811981
let mvarDecl ← mvarId.getDecl
@@ -2013,6 +2013,14 @@ def withoutModifyingMCtx : n α → n α :=
20132013
finally
20142014
modify fun s => { s with cache, mctx }
20152015

2016+
/--
2017+
Ensures that the user names of all local declarations after index `idx` have a macro scope.
2018+
-/
2019+
def _root_.Lean.MVarId.freshenLCtxUserNamesSinceIdx [MonadMCtx n] [MonadLiftT MetaM n] (mvarId : MVarId) (idx : Nat) : n Unit := do
2020+
let some decl := (← getMCtx).findDecl? mvarId | liftMetaM <| throwError m!"unknown metavariable {mvarId.name}"
2021+
let lctx ← freshenUserNamesSinceIdx decl.lctx idx
2022+
modifyMCtx fun mctx => { mctx with decls := mctx.decls.insert mvarId { decl with lctx } }
2023+
20162024
@[inline] private def approxDefEqImp (x : MetaM α) : MetaM α :=
20172025
withConfig (fun config => { config with foApprox := true, ctxApprox := true, quasiPatternApprox := true}) x
20182026

0 commit comments

Comments
 (0)