Skip to content

Commit eb15c08

Browse files
fix: instantiate mvars of indices before instantiating fvars (#4717)
When elaborating the headers of mutual indexed inductive types, mvars have to be synthesized and instantiated before replacing the fvars present there. Otherwise, some fvars present in uninstantiated mvars may be missed and lead to an error later. Closes #3242 (again)
1 parent 72f2e7a commit eb15c08

File tree

2 files changed

+19
-2
lines changed

2 files changed

+19
-2
lines changed

src/Lean/Elab/Inductive.lean

Lines changed: 11 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -771,18 +771,27 @@ private def mkInductiveDecl (vars : Array Expr) (views : Array InductiveView) :
771771
let isUnsafe := view0.modifiers.isUnsafe
772772
withRef view0.ref <| Term.withLevelNames allUserLevelNames do
773773
let rs ← elabHeader views
774+
Term.synthesizeSyntheticMVarsNoPostponing
774775
withInductiveLocalDecls rs fun params indFVars => do
775776
trace[Elab.inductive] "indFVars: {indFVars}"
776777
let mut indTypesArray := #[]
777778
for i in [:views.size] do
778779
let indFVar := indFVars[i]!
779780
Term.addLocalVarInfo views[i]!.declId indFVar
780781
let r := rs[i]!
781-
let type := r.type |>.abstract r.params |>.instantiateRev params
782+
/- At this point, because of `withInductiveLocalDecls`, the only fvars that are in context are the ones related to the first inductive type.
783+
Because of this, we need to replace the fvars present in each inductive type's header of the mutual block with those of the first inductive.
784+
However, some mvars may still be uninstantiated there, and might hide some of the old fvars.
785+
As such we first need to synthesize all possible mvars at this stage, instantiate them in the header types and only
786+
then replace the parameters' fvars in the header type.
787+
788+
See issue #3242 (`https://github.com/leanprover/lean4/issues/3242`)
789+
-/
790+
let type ← instantiateMVars r.type
791+
let type := type.replaceFVars r.params params
782792
let type ← mkForallFVars params type
783793
let ctors ← withExplicitToImplicit params (elabCtors indFVars indFVar params r)
784794
indTypesArray := indTypesArray.push { name := r.view.declName, type, ctors }
785-
Term.synthesizeSyntheticMVarsNoPostponing
786795
let numExplicitParams ← fixedIndicesToParams params.size indTypesArray indFVars
787796
trace[Elab.inductive] "numExplicitParams: {numExplicitParams}"
788797
let indTypes := indTypesArray.toList

tests/lean/run/3242.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -11,3 +11,11 @@ inductive R (F: α → α → Prop) (a : α) : α → Prop where
1111
inductive And₂ (F: α → α → Prop) (a : α) : α → (Nat → α) → Nat → Prop where
1212
| mk (b : α) (f : Nat → α) (n : Nat): R F a (f n) → F (f n) b → And₂ F a b f n
1313
end
14+
15+
structure Salg (n k: Nat) where
16+
D: Type
17+
18+
mutual
19+
inductive Ins (salg: Salg n k) : salg.D → Prop
20+
inductive Out (salg: Salg n k) : salg.D → Prop
21+
end

0 commit comments

Comments
 (0)