You signed in with another tab or window. Reload to refresh your session.You signed out in another tab or window. Reload to refresh your session.You switched accounts on another tab or window. Reload to refresh your session.Dismiss alert
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.
Copy file name to clipboardExpand all lines: tests/lean/inductive1.lean.expected.out
+5-3Lines changed: 5 additions & 3 deletions
Original file line number
Diff line number
Diff line change
@@ -20,10 +20,12 @@ inductive1.lean:80:0-80:27: error: invalid use of 'partial' in inductive declara
20
20
inductive1.lean:81:0-81:33: error: invalid use of 'noncomputable' in inductive declaration
21
21
inductive1.lean:82:2-82:8: error: declaration is not a definition 'T1''
22
22
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
25
25
inductive1.lean:105:7-105:9: error: type expected, got
26
26
(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
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
+
inductiveDescwhere
13
+
| intro
14
+
(name: String)
15
+
(hash: UInt64)
16
+
(params: List Desc)
17
+
: Desc
18
+
deriving Repr
19
+
20
+
defhash_with_name (_name: String) (_params: List Desc): UInt64 := 0-- mock hash function
21
+
22
+
defDesc.intro_func (name: String) (params: List Desc): Desc :=
23
+
Desc.intro
24
+
name
25
+
(hash_with_name name params)
26
+
params
27
+
28
+
inductiveForall {α : 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
+
inductiveIsSmart (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 : Type → Type) : 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
+
inductiveT (α : Type) where
70
+
| mk : NatOf T → T α
71
+
72
+
inductiveT_OK (α : Type) : Type → Type 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
+
inductiveT₂ (α β : Type) : Type
87
+
| mk : NatOf (T₂ α) → T₂ α β
88
+
89
+
abbrev InList : List (Type → Type) → 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
+
inductiveFoo (α : 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
+
inductiveBadIdx (α : 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.
0 commit comments