Skip to content

Commit b93231f

Browse files
authored
feat: improve inductive type parameter error messages (#8338)
This PR improves the error messages displayed in `inductive` declarations when type parameters are invalid or absent. Closes #2195 by improving the relevant error message.
1 parent f40d72e commit b93231f

File tree

4 files changed

+208
-15
lines changed

4 files changed

+208
-15
lines changed

src/Lean/Elab/Inductive.lean

Lines changed: 51 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -138,7 +138,7 @@ private def reorderCtorArgs (ctorType : Expr) : MetaM Expr := do
138138
else
139139
let r ← mkForallFVars (bsPrefix ++ as) type
140140
/- `r` already contains the resulting type.
141-
To be able to produce more better error messages, we copy the first `bsPrefix.size` binder names from `C` to `r`.
141+
To be able to produce better error messages, we copy the first `bsPrefix.size` binder names from `C` to `r`.
142142
This is important when some of constructor parameters were inferred using the auto-bound implicit feature.
143143
For example, in the following declaration.
144144
```
@@ -178,7 +178,8 @@ private def elabCtors (indFVars : Array Expr) (params : Array Expr) (r : ElabHea
178178
match ctorView.type? with
179179
| none =>
180180
if indFamily then
181-
throwError "constructor resulting type must be specified in inductive family declaration"
181+
throwError "Missing resulting type for constructor '{ctorView.declName}': \
182+
Its resulting type must be specified because it is part of an inductive family declaration"
182183
return mkAppN indFVar params
183184
| some ctorType =>
184185
let type ← Term.elabType ctorType
@@ -188,9 +189,9 @@ private def elabCtors (indFVars : Array Expr) (params : Array Expr) (r : ElabHea
188189
let type ← checkParamOccs type
189190
forallTelescopeReducing type fun _ resultingType => do
190191
unless resultingType.getAppFn == indFVar do
191-
throwError "unexpected constructor resulting type{indentExpr resultingType}"
192+
throwUnexpectedResultingTypeMismatch resultingType indFVar ctorView.declName ctorType
192193
unless (← isType resultingType) do
193-
throwError "unexpected constructor resulting type, type expected{indentExpr resultingType}"
194+
throwUnexpectedResultingTypeNotType resultingType ctorView.declName ctorType
194195
return type
195196
let type ← elabCtorType
196197
Term.synthesizeSyntheticMVarsNoPostponing
@@ -229,23 +230,62 @@ private def elabCtors (indFVars : Array Expr) (params : Array Expr) (r : ElabHea
229230
trace[Elab.inductive] "{ctorView.declName} : {type}"
230231
return { name := ctorView.declName, type }
231232
where
233+
/--
234+
Ensures that the arguments to recursive occurrences of the type family being defined match the
235+
parameters from the inductive definition.
236+
-/
232237
checkParamOccs (ctorType : Expr) : MetaM Expr :=
233-
let visit (e : Expr) : MetaM TransformStep := do
238+
let visit (e : Expr) : StateT (List Expr) MetaM TransformStep := do
234239
let f := e.getAppFn
235240
if indFVars.contains f then
236241
let mut args := e.getAppArgs
237-
unless args.size ≥ params.size do
238-
throwError "unexpected inductive type occurrence{indentExpr e}"
239-
for h : i in [:params.size] do
240-
let param := params[i]
242+
-- Prefer throwing an "argument mismatch" error rather than a "missing parameter" one
243+
for i in [:min args.size params.size] do
244+
let param := params[i]!
241245
let arg := args[i]!
242246
unless (← isDefEq param arg) do
243-
throwError "inductive datatype parameter mismatch{indentExpr arg}\nexpected{indentExpr param}"
247+
let (arg, param) ← addPPExplicitToExposeDiff arg param
248+
let msg := m!"Mismatched inductive type parameter in{indentExpr e}\nThe provided argument\
249+
{indentExpr arg}\nis not definitionally equal to the expected parameter{indentExpr param}"
250+
let noteMsg := m!"The value of parameter '{param}' must be fixed throughout the inductive \
251+
declaration. Consider making this parameter an index if it must vary."
252+
throwError msg ++ .note noteMsg
244253
args := args.set! i param
254+
unless args.size ≥ params.size do
255+
let expected := mkAppN f params
256+
let containingExprMsg := (← get).head?.map toMessageData |>.getD m!"<missing>"
257+
let msg :=
258+
m!"Missing parameter(s) in occurrence of inductive type: In the expression{indentD containingExprMsg}\n\
259+
found{indentExpr e}\nbut expected all parameters to be specified:{indentExpr expected}"
260+
let noteMsg :=
261+
m!"All occurrences of an inductive type in the types of its constructors must specify its \
262+
fixed parameters. Only indices can be omitted in a partial application of the type constructor."
263+
throwError msg ++ .note noteMsg
245264
return TransformStep.done (mkAppN f args)
246265
else
266+
modify fun es => e :: es
247267
return .continue
248-
transform ctorType (pre := visit)
268+
let popContainingExpr (e : Expr) : StateT (List Expr) MetaM TransformStep := do
269+
modify fun es => es.drop 1
270+
return .done e
271+
transform ctorType (pre := visit) (post := popContainingExpr) |>.run' [ctorType]
272+
273+
throwUnexpectedResultingTypeMismatch (resultingType indFVar : Expr) (declName : Name) (ctorType : Syntax) := do
274+
let lazyAppMsg := MessageData.ofLazyM do
275+
if let some fvarId := indFVar.fvarId? then
276+
if let some decl := (← getLCtx).find? fvarId then
277+
if (← whnfD decl.type).isForall then
278+
return m!" an application of"
279+
return m!""
280+
throwErrorAt ctorType "Unexpected resulting type for constructor '{declName}': \
281+
Expected{lazyAppMsg}{indentExpr indFVar}\nbut found{indentExpr resultingType}"
282+
283+
throwUnexpectedResultingTypeNotType (resultingType : Expr) (declName : Name) (ctorType : Syntax) := do
284+
let lazyMsg := MessageData.ofLazyM do
285+
let resultingTypeType ← inferType resultingType
286+
return indentExpr resultingTypeType
287+
throwErrorAt ctorType "Unexpected resulting type for constructor '{declName}': \
288+
Expected a type, but found{indentExpr resultingType}\nof type{lazyMsg}"
249289

250290
@[builtin_inductive_elab Lean.Parser.Command.inductive, builtin_inductive_elab Lean.Parser.Command.classInductive]
251291
def elabInductiveCommand : InductiveElabDescr where

src/Lean/Elab/MutualInductive.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -219,7 +219,7 @@ private def checkUnsafe (rs : Array PreElabHeaderResult) : TermElabM Unit := do
219219
let isUnsafe := rs[0]!.view.modifiers.isUnsafe
220220
for r in rs do
221221
unless r.view.modifiers.isUnsafe == isUnsafe do
222-
throwErrorAt r.view.ref "invalid inductive type, cannot mix unsafe and safe declarations in a mutually inductive datatypes"
222+
throwErrorAt r.view.ref "invalid inductive type, cannot mix unsafe and safe declarations in mutually inductive datatypes"
223223

224224
private def checkClass (rs : Array PreElabHeaderResult) : TermElabM Unit := do
225225
if rs.size > 1 then

tests/lean/inductive1.lean.expected.out

Lines changed: 5 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -20,10 +20,12 @@ inductive1.lean:80:0-80:27: error: invalid use of 'partial' in inductive declara
2020
inductive1.lean:81:0-81:33: error: invalid use of 'noncomputable' in inductive declaration
2121
inductive1.lean:82:2-82:8: error: declaration is not a definition 'T1''
2222
inductive1.lean:85:0-85:17: error: invalid 'private' constructor in a 'private' inductive datatype
23-
inductive1.lean:93:7-93:26: error: invalid inductive type, cannot mix unsafe and safe declarations in a mutually inductive datatypes
24-
inductive1.lean:100:0-100:4: error: constructor resulting type must be specified in inductive family declaration
23+
inductive1.lean:93:7-93:26: error: invalid inductive type, cannot mix unsafe and safe declarations in mutually inductive datatypes
24+
inductive1.lean:100:0-100:4: error: Missing resulting type for constructor 'T1.z2': Its resulting type must be specified because it is part of an inductive family declaration
2525
inductive1.lean:105:7-105:9: error: type expected, got
2626
(T1 : Nat → Type)
27-
inductive1.lean:108:0-108:10: error: unexpected constructor resulting type
27+
inductive1.lean:108:7-108:10: error: Unexpected resulting type for constructor 'T1.z1': Expected an application of
28+
T1
29+
but found
2830
Nat
2931
inductive1.lean:118:7-118:11: error: unknown identifier 'cons'
Lines changed: 151 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,151 @@
1+
/-! # Inductive parameter mismatch error messages
2+
3+
Tests that appropriate error messages are shown when the fixed parameters of an inductive type
4+
constructor are omitted or incorrect in an `inductive` declaration.
5+
6+
A previous version of one such error message was noted to be confusing in #2195:
7+
https://github.com/leanprover/lean4/issues/2195
8+
-/
9+
10+
/-! ## Example from Issue #2195 -/
11+
12+
inductive Desc where
13+
| intro
14+
(name: String)
15+
(hash: UInt64)
16+
(params: List Desc)
17+
: Desc
18+
deriving Repr
19+
20+
def hash_with_name (_name: String) (_params: List Desc): UInt64 := 0 -- mock hash function
21+
22+
def Desc.intro_func (name: String) (params: List Desc): Desc :=
23+
Desc.intro
24+
name
25+
(hash_with_name name params)
26+
params
27+
28+
inductive Forall {α : Type u} (p : α → Prop) : List α → Prop
29+
| nil : Forall p ([] : List α)
30+
| cons : ∀ {x xs}, p x → Forall p xs → Forall p (x :: xs)
31+
32+
/--
33+
error: Missing parameter(s) in occurrence of inductive type: In the expression
34+
Forall IsSmart params
35+
found
36+
IsSmart
37+
but expected all parameters to be specified:
38+
IsSmart d
39+
40+
Note: All occurrences of an inductive type in the types of its constructors must specify its fixed parameters. Only indices can be omitted in a partial application of the type constructor.
41+
-/
42+
#guard_msgs in
43+
inductive IsSmart (d: Desc): Prop
44+
| isSmart: ∀
45+
(name: String)
46+
(params: List Desc)
47+
(hash: UInt64)
48+
(reader: Bool),
49+
d = Desc.intro name hash params
50+
→ hash = hash_with_name name params
51+
→ Forall IsSmart params
52+
→ IsSmart d
53+
54+
55+
/-! ## "Missing parameter" error -/
56+
57+
abbrev NatOf (F : TypeType) : Type := F Nat
58+
/--
59+
error: Missing parameter(s) in occurrence of inductive type: In the expression
60+
NatOf T
61+
found
62+
T
63+
but expected all parameters to be specified:
64+
T α
65+
66+
Note: All occurrences of an inductive type in the types of its constructors must specify its fixed parameters. Only indices can be omitted in a partial application of the type constructor.
67+
-/
68+
#guard_msgs in
69+
inductive T (α : Type) where
70+
| mk : NatOf T → T α
71+
72+
inductive T_OK (α : Type) : TypeType where
73+
| mk : NatOf (T_OK α) → T_OK α Nat
74+
75+
/--
76+
error: Missing parameter(s) in occurrence of inductive type: In the expression
77+
NatOf (T₂ α)
78+
found
79+
T₂ α
80+
but expected all parameters to be specified:
81+
T₂ α β
82+
83+
Note: All occurrences of an inductive type in the types of its constructors must specify its fixed parameters. Only indices can be omitted in a partial application of the type constructor.
84+
-/
85+
#guard_msgs in
86+
inductive T₂ (α β : Type) : Type
87+
| mk : NatOf (T₂ α) → T₂ α β
88+
89+
abbrev InList : List (TypeType) → Type :=
90+
fun _ => Nat
91+
/--
92+
error: Missing parameter(s) in occurrence of inductive type: In the expression
93+
[Foo]
94+
found
95+
Foo
96+
but expected all parameters to be specified:
97+
Foo α
98+
99+
Note: All occurrences of an inductive type in the types of its constructors must specify its fixed parameters. Only indices can be omitted in a partial application of the type constructor.
100+
-/
101+
#guard_msgs in
102+
inductive Foo (α : Type) : Type
103+
| mk : InList [Foo] → Foo α
104+
105+
106+
/-! ## "Mismatched parameter" error -/
107+
108+
/--
109+
error: Mismatched inductive type parameter in
110+
BadIdx α 0
111+
The provided argument
112+
0
113+
is not definitionally equal to the expected parameter
114+
n
115+
116+
Note: The value of parameter 'n' must be fixed throughout the inductive declaration. Consider making this parameter an index if it must vary.
117+
-/
118+
#guard_msgs in
119+
inductive BadIdx (α : Type) (n : Nat) : Type
120+
| mk : BadIdx α 0
121+
122+
/--
123+
error: Mismatched inductive type parameter in
124+
BadIdx' α 0 k
125+
The provided argument
126+
0
127+
is not definitionally equal to the expected parameter
128+
n
129+
130+
Note: The value of parameter 'n' must be fixed throughout the inductive declaration. Consider making this parameter an index if it must vary.
131+
-/
132+
#guard_msgs in
133+
inductive BadIdx' (α : Type) (n k : Nat) : Type
134+
| mk : BadIdx' α 0 k
135+
136+
137+
/-! ## "Mismatched parameter" preempts "missing parameter" -/
138+
139+
/--
140+
error: Mismatched inductive type parameter in
141+
Bar Nat
142+
The provided argument
143+
Nat
144+
is not definitionally equal to the expected parameter
145+
α
146+
147+
Note: The value of parameter 'α' must be fixed throughout the inductive declaration. Consider making this parameter an index if it must vary.
148+
-/
149+
#guard_msgs in
150+
inductive Bar (α β : Type) : Type
151+
| mk : NatOf (Bar Nat) → Bar α β

0 commit comments

Comments
 (0)