Skip to content

Commit f88e503

Browse files
authored
feat: @[suggest_for] annotations for prompting easy-to-miss names (#11554)
This PR adds `@[suggest_for]` annotations to Lean, allowing lean to provide corrections for `.every` or `.some` methods in place of `.all` or `.any` methods for most default-imported types (arrays, lists, strings, substrings, and subarrays, and vectors). Due to the need for stage0 updates for new annotations, the `suggest_for` annotation itself was introduced in previous PRs: #11367, #11529, and #11590. ## Example ``` example := "abc".every (! ·.isWhitespace) ``` Error message: ``` Invalid field `every`: The environment does not contain `String.every`, so it is not possible to project the field `every` from an expression "abc" of type `String` Hint: Perhaps you meant `String.all` in place of `String.every`: .e̵v̵e̵r̵y̵a̲l̲l̲ ``` (the hint is added by this PR) ## Additional changes Adds suggestions that are not currently active but that can be used to generate autocompletion suggestions in the reference manual: - `Either` -> `Except` and `Sum` - `Exception` -> `Except` - `ℕ` -> `Nat` - `Nullable` -> `Option` - `Maybe` -> `Option` - `Optional` -> `Option` - `Result` -> `Except`
1 parent 2e4b079 commit f88e503

File tree

11 files changed

+141
-11
lines changed

11 files changed

+141
-11
lines changed

src/Init/Core.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -201,6 +201,7 @@ An element of `α ⊕ β` is either an `a : α` wrapped in `Sum.inl` or a `b :
201201
indication of which of the two types was chosen. The union of a singleton set with itself contains
202202
one element, while `Unit ⊕ Unit` contains distinct values `inl ()` and `inr ()`.
203203
-/
204+
@[suggest_for Either]
204205
inductive Sum (α : Type u) (β : Type v) where
205206
/-- Left injection into the sum type `α ⊕ β`. -/
206207
| inl (val : α) : Sum α β

src/Init/Data/Array/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1348,7 +1348,7 @@ Examples:
13481348
* `#[2, 4, 5, 6].any (· % 2 = 0) = true`
13491349
* `#[2, 4, 5, 6].any (· % 2 = 1) = true`
13501350
-/
1351-
@[inline, expose]
1351+
@[inline, expose, suggest_for Array.some]
13521352
def any (as : Array α) (p : α → Bool) (start := 0) (stop := as.size) : Bool :=
13531353
Id.run <| as.anyM (pure <| p ·) start stop
13541354

@@ -1366,7 +1366,7 @@ Examples:
13661366
* `#[2, 4, 6].all (· % 2 = 0) = true`
13671367
* `#[2, 4, 5, 6].all (· % 2 = 0) = false`
13681368
-/
1369-
@[inline]
1369+
@[inline, suggest_for Array.every]
13701370
def all (as : Array α) (p : α → Bool) (start := 0) (stop := as.size) : Bool :=
13711371
Id.run <| as.allM (pure <| p ·) start stop
13721372

src/Init/Data/Array/Subarray.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -280,7 +280,7 @@ Checks whether any of the elements in a subarray satisfy a Boolean predicate.
280280
The elements are tested starting at the lowest index and moving up. The search terminates as soon as
281281
an element that satisfies the predicate is found.
282282
-/
283-
@[inline]
283+
@[inline, suggest_for Subarray.some]
284284
def any {α : Type u} (p : α → Bool) (as : Subarray α) : Bool :=
285285
Id.run <| as.anyM (pure <| p ·)
286286

@@ -290,7 +290,7 @@ Checks whether all of the elements in a subarray satisfy a Boolean predicate.
290290
The elements are tested starting at the lowest index and moving up. The search terminates as soon as
291291
an element that does not satisfy the predicate is found.
292292
-/
293-
@[inline]
293+
@[inline, suggest_for Subarray.every]
294294
def all {α : Type u} (p : α → Bool) (as : Subarray α) : Bool :=
295295
Id.run <| as.allM (pure <| p ·)
296296

src/Init/Data/List/Basic.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1847,6 +1847,7 @@ Examples:
18471847
* `[2, 4, 5, 6].any (· % 2 = 0) = true`
18481848
* `[2, 4, 5, 6].any (· % 2 = 1) = true`
18491849
-/
1850+
@[suggest_for List.some]
18501851
def any : (l : List α) → (p : α → Bool) → Bool
18511852
| [], _ => false
18521853
| h :: t, p => p h || any t p
@@ -1866,6 +1867,7 @@ Examples:
18661867
* `[2, 4, 6].all (· % 2 = 0) = true`
18671868
* `[2, 4, 5, 6].all (· % 2 = 0) = false`
18681869
-/
1870+
@[suggest_for List.every]
18691871
def all : List α → (α → Bool) → Bool
18701872
| [], _ => true
18711873
| h :: t, p => p h && all t p

src/Init/Data/String/Search.lean

Lines changed: 3 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -324,7 +324,8 @@ Examples:
324324
* {lean}`"tea".contains (fun (c : Char) => c == 'X') = false`
325325
* {lean}`"coffee tea water".contains "tea" = true`
326326
-/
327-
@[inline] def contains (s : String) (pat : ρ) [ToForwardSearcher pat σ] : Bool :=
327+
@[inline, suggest_for String.some]
328+
def contains (s : String) (pat : ρ) [ToForwardSearcher pat σ] : Bool :=
328329
s.toSlice.contains pat
329330

330331
@[export lean_string_contains]
@@ -352,7 +353,7 @@ Examples:
352353
* {lean}`"aaaaaa".all "aa" = true`
353354
* {lean}`"aaaaaaa".all "aa" = false`
354355
-/
355-
@[inline] def all (s : String) (pat : ρ) [ForwardPattern pat] : Bool :=
356+
@[inline, suggest_for String.every] def all (s : String) (pat : ρ) [ForwardPattern pat] : Bool :=
356357
s.toSlice.all pat
357358

358359
/--

src/Init/Data/String/Slice.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -505,7 +505,7 @@ Examples:
505505
* {lean}`"tea".toSlice.contains (fun (c : Char) => c == 'X') = false`
506506
* {lean}`"coffee tea water".toSlice.contains "tea" = true`
507507
-/
508-
@[specialize pat]
508+
@[specialize pat, suggest_for String.Slice.some]
509509
def contains (s : Slice) (pat : ρ) [ToForwardSearcher pat σ] : Bool :=
510510
let searcher := ToForwardSearcher.toSearcher pat s
511511
searcher.any (· matches .matched ..)

src/Init/Data/String/Substring.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -274,15 +274,15 @@ Checks whether the Boolean predicate `p` returns `true` for any character in a s
274274
275275
Short-circuits at the first character for which `p` returns `true`.
276276
-/
277-
@[inline] def any (s : Substring.Raw) (p : Char → Bool) : Bool :=
277+
@[inline, suggest_for Substring.Raw.some] def any (s : Substring.Raw) (p : Char → Bool) : Bool :=
278278
s.toSlice?.get!.any p
279279

280280
/--
281281
Checks whether the Boolean predicate `p` returns `true` for every character in a substring.
282282
283283
Short-circuits at the first character for which `p` returns `false`.
284284
-/
285-
@[inline] def all (s : Substring.Raw) (p : Char → Bool) : Bool :=
285+
@[inline, suggest_for Substring.Raw.every] def all (s : Substring.Raw) (p : Char → Bool) : Bool :=
286286
s.toSlice?.get!.all p
287287

288288
@[export lean_substring_all]

src/Init/Data/Vector/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -474,11 +474,11 @@ to avoid having to have the predicate live in `p : α → m (ULift Bool)`.
474474
xs.toArray.allM p
475475

476476
/-- Returns `true` if `p` returns `true` for any element of the vector. -/
477-
@[inline, expose] def any (xs : Vector α n) (p : α → Bool) : Bool :=
477+
@[inline, expose, suggest_for Vector.some] def any (xs : Vector α n) (p : α → Bool) : Bool :=
478478
xs.toArray.any p
479479

480480
/-- Returns `true` if `p` returns `true` for all elements of the vector. -/
481-
@[inline, expose] def all (xs : Vector α n) (p : α → Bool) : Bool :=
481+
@[inline, expose, suggest_for Vector.every] def all (xs : Vector α n) (p : α → Bool) : Bool :=
482482
xs.toArray.all p
483483

484484
/-- Count the number of elements of a vector that satisfy the predicate `p`. -/

src/Init/Prelude.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1197,6 +1197,7 @@ This type is special-cased by both the kernel and the compiler, and overridden w
11971197
implementation. Both use a fast arbitrary-precision arithmetic library (usually
11981198
[GMP](https://gmplib.org/)); at runtime, `Nat` values that are sufficiently small are unboxed.
11991199
-/
1200+
@[suggest_for ℕ]
12001201
inductive Nat where
12011202
/--
12021203
Zero, the smallest natural number.
@@ -2856,6 +2857,7 @@ Optional values, which are either `some` around a value from the underlying type
28562857
`Option` can represent nullable types or computations that might fail. In the codomain of a function
28572858
type, it can also represent partiality.
28582859
-/
2860+
@[suggest_for Maybe, suggest_for Optional, suggest_for Nullable]
28592861
inductive Option (α : Type u) where
28602862
/-- No value. -/
28612863
| none : Option α
@@ -3939,6 +3941,7 @@ value of type `α`.
39393941
the `pure` operation is `Except.ok` and the `bind` operation returns the first encountered
39403942
`Except.error`.
39413943
-/
3944+
@[suggest_for Result, suggest_for Exception, suggest_for Either]
39423945
inductive Except (ε : Type u) (α : Type v) where
39433946
/-- A failure value of type `ε` -/
39443947
| error : ε → Except ε α

tests/lean/run/idSuggestEvery.lean

Lines changed: 110 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,110 @@
1+
-- test every/all replacements in default imports
2+
3+
/-
4+
Expected replacements for `Subarray.any` and `Subarray.all` do not work as suggestion annotations
5+
when using generalized field notation: Subarray is an abbreviation and so the underlying `Std.Slice`
6+
type, which does not have a corresponding `.any`/`.all` function.
7+
-/
8+
9+
/--
10+
error: Invalid field `every`: The environment does not contain `Std.Slice.every`, so it is not possible to project the field `every` from an expression
11+
x
12+
of type
13+
Std.Slice (Std.Slice.Internal.SubarrayData Nat)
14+
-/
15+
#guard_msgs in example (x : Subarray Nat) := x.every fun _ => true
16+
#guard_msgs in example (x : Subarray Nat) := x.all fun _ => true
17+
18+
/--
19+
error: Unknown constant `Subarray.every`
20+
21+
Hint: Perhaps you meant `Subarray.all` in place of `Subarray.every`:
22+
[apply] `Subarray.all`
23+
-/
24+
#guard_msgs in example := (@Subarray.every Nat (fun _ => true) ·)
25+
#guard_msgs in example := (@Subarray.all Nat (fun _ => true) ·)
26+
27+
/--
28+
error: Unknown constant `Subarray.some`
29+
30+
Hint: Perhaps you meant `Subarray.any` in place of `Subarray.some`:
31+
[apply] `Subarray.any`
32+
-/
33+
#guard_msgs in example := (@Subarray.some Nat (fun _ => true) ·)
34+
#guard_msgs in example := (@Subarray.any Nat (fun _ => true) ·)
35+
36+
/--
37+
error: Invalid field `every`: The environment does not contain `String.every`, so it is not possible to project the field `every` from an expression
38+
x
39+
of type `String`
40+
41+
Hint: Perhaps you meant `String.all` in place of `String.every`:
42+
.e̵v̵e̵r̵y̵a̲l̲l̲
43+
-/
44+
#guard_msgs in example (x : String) := x.every fun _ => true
45+
#guard_msgs in example (x : String) := x.all fun _ => true
46+
/--
47+
error: Invalid field `some`: The environment does not contain `String.some`, so it is not possible to project the field `some` from an expression
48+
x
49+
of type `String`
50+
51+
Hint: Perhaps you meant `String.contains` in place of `String.some`:
52+
.s̵o̵m̵e̵c̲o̲n̲t̲a̲i̲n̲s̲
53+
-/
54+
#guard_msgs in example (x : String) := x.some fun _ => true
55+
#guard_msgs in example (x : String) := x.contains fun _ => true
56+
57+
58+
/--
59+
error: Invalid field `every`: The environment does not contain `Array.every`, so it is not possible to project the field `every` from an expression
60+
x
61+
of type `Array Nat`
62+
63+
Hint: Perhaps you meant `Array.all` in place of `Array.every`:
64+
.e̵v̵e̵r̵y̵a̲l̲l̲
65+
-/
66+
#guard_msgs in example (x : Array Nat) := x.every fun _ => true
67+
#guard_msgs in example (x : Array Nat) := x.all fun _ => true
68+
/--
69+
error: Invalid field `some`: The environment does not contain `Array.some`, so it is not possible to project the field `some` from an expression
70+
x
71+
of type `Array Nat`
72+
73+
Hint: Perhaps you meant `Array.any` in place of `Array.some`:
74+
.s̵o̵m̵e̵a̲n̲y̲
75+
-/
76+
#guard_msgs in example (x : Array Nat) := x.some fun _ => true
77+
#guard_msgs in example (x : Array Nat) := x.all fun _ => true
78+
79+
/--
80+
error: Invalid field `every`: The environment does not contain `List.every`, so it is not possible to project the field `every` from an expression
81+
x
82+
of type `List Nat`
83+
84+
Hint: Perhaps you meant `List.all` in place of `List.every`:
85+
.e̵v̵e̵r̵y̵a̲l̲l̲
86+
-/
87+
#guard_msgs in example (x : List Nat) := x.every fun _ => true
88+
#guard_msgs in example (x : List Nat) := x.all fun _ => true
89+
/--
90+
error: Invalid field `some`: The environment does not contain `List.some`, so it is not possible to project the field `some` from an expression
91+
x
92+
of type `List Nat`
93+
94+
Hint: Perhaps you meant `List.any` in place of `List.some`:
95+
.s̵o̵m̵e̵a̲n̲y̲
96+
-/
97+
#guard_msgs in example (x : List Nat) := x.some fun _ => true
98+
#guard_msgs in example (x : List Nat) := x.all fun _ => true
99+
100+
/--
101+
@ +1:17...22
102+
error: Invalid field `every`: The environment does not contain `List.every`, so it is not possible to project the field `every` from an expression
103+
[1, 2, 3]
104+
of type `List ?m.10`
105+
106+
Hint: Perhaps you meant `List.all` in place of `List.every`:
107+
.e̵v̵e̵r̵y̵a̲l̲l̲
108+
-/
109+
#guard_msgs (positions := true) in
110+
#check [1, 2, 3].every

0 commit comments

Comments
 (0)