Skip to content

Commit ffd568d

Browse files
committed
cleanup of several aspects of identifier suggestion
1 parent 3c100ad commit ffd568d

14 files changed

+128
-68
lines changed

src/Lean/Elab/App.lean

Lines changed: 24 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -1314,6 +1314,16 @@ where
13141314
| none => 0
13151315
| some (_, p2) => prodArity p2 + 1
13161316

1317+
/--
1318+
For a fieldname `fieldName`, locate all the environment constants with that name
1319+
-/
1320+
private def reverseFieldLookup (env : Environment) (fieldName : String) :=
1321+
env.constants.fold (init := #[]) (fun accum name _ =>
1322+
match name with
1323+
| .str _ s => if s = fieldName && !name.isInternal then accum.push name else accum
1324+
| _ => accum)
1325+
|>.qsort (lt := Name.lt)
1326+
13171327
private def resolveLValAux (e : Expr) (eType : Expr) (lval : LVal) : TermElabM LValResolution := do
13181328
match eType.getAppFn, lval with
13191329
| .const structName _, LVal.fieldIdx ref idx =>
@@ -1379,11 +1389,7 @@ private def resolveLValAux (e : Expr) (eType : Expr) (lval : LVal) : TermElabM L
13791389
has function type{inlineExprTrailing eType}"
13801390

13811391
| .mvar .., .fieldName _ fieldName _ _ =>
1382-
let possibleConstants := (← getEnv).constants.fold (fun accum name _ =>
1383-
match name with
1384-
| .str _ s => if s = fieldName && !name.isInternal then accum.push name else accum
1385-
| _ => accum) #[]
1386-
let hint := match possibleConstants with
1392+
let hint := match reverseFieldLookup (← getEnv) fieldName with
13871393
| #[] => MessageData.nil
13881394
| #[opt] => .hint' m!"Consider replacing the field projection `.{fieldName}` with a call to the function `{.ofConstName opt}`."
13891395
| opts => .hint' m!"Consider replacing the field projection with a call to one of the following:\
@@ -1428,16 +1434,15 @@ where
14281434
(suggestions.map fun suggestion => {
14291435
preInfo? := .some s!"{e}.",
14301436
suggestion := suggestion.getString!,
1431-
toCodeActionTitle? := .some (s!"Suggested replacement: {e}.{·}"),
1437+
toCodeActionTitle? := .some (s!"Change to {e}.{·}"),
14321438
diffGranularity := .all,
14331439
})
14341440
else
14351441
MessageData.hint (ref? := ref)
14361442
m!"Perhaps you meant one of these in place of `{fullName}`:"
14371443
(suggestions.map fun suggestion => {
14381444
suggestion := suggestion.getString!,
1439-
toCodeActionTitle? := .some (s!"Suggested replacement: {e}.{·}"),
1440-
diffGranularity := .all,
1445+
toCodeActionTitle? := .some (s!"Change to {e}.{·}"),
14411446
messageData? := .some m!"`{.ofConstName suggestion}`: {e}.{suggestion.getString!}",
14421447
})
14431448

@@ -1765,8 +1770,17 @@ private partial def resolveDottedIdentFn (idRef : Syntax) (id : Name) (expectedT
17651770
withForallBody expectedType fun resultType => do
17661771
go resultType expectedType #[]
17671772
where
1768-
throwNoExpectedType :=
1769-
throwNamedError lean.invalidDottedIdent "Invalid dotted identifier notation: The expected type of `.{id}` could not be determined"
1773+
throwNoExpectedType := do
1774+
let hint ← match reverseFieldLookup (← getEnv) (id.getString!) with
1775+
| #[] => pure MessageData.nil
1776+
| suggestions =>
1777+
let oneOfThese := if h : suggestions.size = 1 then m!"`{.ofConstName suggestions[0]}" else m!"one of these"
1778+
m!"Using {oneOfThese} would be unambiguous:".hint (suggestions.map fun suggestion => {
1779+
suggestion := mkIdent suggestion
1780+
toCodeActionTitle? := .some (s!"Change to {·}")
1781+
messageData? := .some m!"`{.ofConstName suggestion}`",
1782+
})
1783+
throwNamedError lean.invalidDottedIdent (m!"Invalid dotted identifier notation: The expected type of `.{id}` could not be determined" ++ hint)
17701784
/-- A weak version of forallTelescopeReducing that only uses whnfCore, to avoid unfolding definitions except by `unfoldDefinition?` below. -/
17711785
withForallBody {α} (type : Expr) (k : Expr → TermElabM α) : TermElabM α := do
17721786
let type ← whnfCoreUnfoldingAnnotations type

src/Lean/Elab/PatternVar.lean

Lines changed: 7 additions & 29 deletions
Original file line numberDiff line numberDiff line change
@@ -67,35 +67,13 @@ private def throwCtorExpected {α} (ident : Option Syntax) : M α := do
6767

6868
if candidates.size = 0 then
6969
throwError message
70-
else if h : candidates.size = 1 then
71-
throwError message ++ .hint' m!"`{candidates[0]}` is similar"
72-
else
73-
let sorted := candidates.qsort (·.toString < ·.toString)
74-
let diff :=
75-
if candidates.size > 10 then [m!" (or {candidates.size - 10} others)"]
76-
else []
77-
let suggestions : MessageData := .group <|
78-
.joinSep ((sorted.extract 0 10 |>.toList |>.map (showName env)) ++ diff)
79-
("," ++ Format.line)
80-
throwError message ++ .group (.hint' ("These are similar:" ++ .nestD (Format.line ++ suggestions)))
81-
where
82-
-- Create some `MessageData` for a name that shows it without an `@`, but with the metadata that
83-
-- makes infoview hovers and the like work. This technique only works because the names are known
84-
-- to be global constants, so we don't need the local context.
85-
showName (env : Environment) (n : Name) : MessageData :=
86-
let params :=
87-
env.constants.find?' n |>.map (·.levelParams.map Level.param) |>.getD []
88-
.ofFormatWithInfos {
89-
fmt := "'" ++ .tag 0 (format n) ++ "'",
90-
infos :=
91-
.ofList [(0, .ofTermInfo {
92-
lctx := .empty,
93-
expr := .const n params,
94-
stx := .ident .none (toString n).toRawSubstring n [.decl n []],
95-
elaborator := `Delab,
96-
expectedType? := none
97-
})] _
98-
}
70+
let oneOfThese := if h : candidates.size = 1 then m!"`{candidates[0]}`" else m!"one of these"
71+
let hint ← m!"Using {oneOfThese} would be valid:".hint (ref? := idStx) (candidates.map fun candidate => {
72+
suggestion := mkIdent candidate
73+
toCodeActionTitle? := .some (s!"Change to {·}")
74+
messageData? := .some m!"`{.ofConstName candidate}`",
75+
})
76+
throwError message ++ hint
9977

10078
private def throwInvalidPattern {α} : M α :=
10179
throwError "Invalid pattern"

src/Lean/IdentifierSuggestion.lean

Lines changed: 2 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -90,8 +90,7 @@ public def throwUnknownConstantWithSuggestions (constName : Name) : CoreM α :=
9090
let modified := modifySuggestion suggestion
9191
{
9292
suggestion := s!"{modified}",
93-
toCodeActionTitle? := .some (s!"Suggested replacement: {·}"),
94-
diffGranularity := .all,
95-
-- messageData? := .some m!"replace `{.ofName rawId}` with `{.ofName modified}`",
93+
toCodeActionTitle? := .some (s!"Change to {·}"),
94+
messageData? := .some m!"`{.ofConstName suggestion}`",
9695
}) ref
9796
throwUnknownIdentifierAt (declHint := constName) ref (m!"Unknown constant `{.ofConstName constName}`" ++ hint)

tests/lean/10488.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -51,3 +51,10 @@ structure Foo where
5151
def bar : Foo := {}
5252
#check bar.x.2.snd
5353
#eval { bar with x.2.snd := 1 }
54+
55+
inductive Nope where
56+
| succ : Nope -> Nope -> Nope
57+
| leaf : Nope
58+
59+
example := (match · with
60+
| succ x y => 4)

tests/lean/10488.lean.expected.out

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -30,6 +30,12 @@ foo 31 : ?m → Unit
3030
foo 31 : ?m → Unit
3131
10488.lean:33:15-33:16: error: unexpected token '('; expected command
3232
10488.lean:34:14-34:19: error(lean.invalidDottedIdent): Invalid dotted identifier notation: The expected type of `.succ` could not be determined
33+
34+
Hint: Using one of these would be unambiguous:
35+
[apply] `succ`
36+
[apply] `Fin.succ`
37+
[apply] `Nat.succ`
38+
[apply] `Std.PRange.succ`
3339
foo 31 sorry : Unit
3440
foo 31. succ : Unit
3541
11 : Nat
@@ -41,3 +47,8 @@ foo 31. succ : Unit
4147
10488.lean:41:13-41:17: error: unexpected identifier; expected ')', ',' or ':'
4248
bar.x.snd.snd : Nat
4349
{ x := (2, 4, 1) }
50+
10488.lean:60:4-60:12: error: Invalid pattern: Expected a constructor or constant marked with `[match_pattern]`
51+
52+
Hint: Using one of these would be valid:
53+
[apply] `Nat.succ`
54+
[apply] `Nope.succ`

tests/lean/1719.lean.expected.out

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1 +1,8 @@
11
1719.lean:2:8-2:12: error(lean.invalidDottedIdent): Invalid dotted identifier notation: The expected type of `.inl` could not be determined
2+
3+
Hint: Using one of these would be unambiguous:
4+
[apply] `Or.inl`
5+
[apply] `PSum.inl`
6+
[apply] `Sum.inl`
7+
[apply] `Sum.Lex.inl`
8+
[apply] `Sum.LiftRel.inl`

tests/lean/coeM.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,10 @@ It was checking that `Eq (some true)` and `Eq _` were defeq monads. The defeq ch
2626
-/
2727
/--
2828
error: Invalid dotted identifier notation: The expected type of `.some` could not be determined
29+
30+
Hint: Using one of these would be unambiguous:
31+
[apply] `some`
32+
[apply] `Option.Rel.some`
2933
-/
3034
#guard_msgs in
3135
example : some true = (some true).map id := by

tests/lean/emptyTypeAscription.lean.expected.out

Lines changed: 20 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,25 @@
11
emptyTypeAscription.lean:3:19-3:24: error(lean.invalidDottedIdent): Invalid dotted identifier notation: The expected type of `.zero` could not be determined
2+
3+
Hint: Using one of these would be unambiguous:
4+
[apply] `BitVec.zero`
5+
[apply] `Dyadic.zero`
6+
[apply] `Nat.zero`
7+
[apply] `Vector.zero`
8+
[apply] `Zero.zero`
9+
[apply] `System.Uri.UriEscape.zero`
10+
[apply] `Lean.Grind.IntModule.OfNatModule.zero`
11+
[apply] `Lean.Grind.Linarith.Expr.zero`
212
emptyTypeAscription.lean:4:29-4:34: error(lean.invalidDottedIdent): Invalid dotted identifier notation: The expected type of `.zero` could not be determined
13+
14+
Hint: Using one of these would be unambiguous:
15+
[apply] `BitVec.zero`
16+
[apply] `Dyadic.zero`
17+
[apply] `Nat.zero`
18+
[apply] `Vector.zero`
19+
[apply] `Zero.zero`
20+
[apply] `System.Uri.UriEscape.zero`
21+
[apply] `Lean.Grind.IntModule.OfNatModule.zero`
22+
[apply] `Lean.Grind.Linarith.Expr.zero`
323
emptyTypeAscription.lean:9:63-9:70: error: invalid `▸` notation, expected result type of cast is
424
w = add x y
525
however, the equality

tests/lean/run/change.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -139,6 +139,10 @@ Fails, type hint can't hint enough since `.some _` is postponed.
139139
-/
140140
/--
141141
error: Invalid dotted identifier notation: The expected type of `.some` could not be determined
142+
143+
Hint: Using one of these would be unambiguous:
144+
[apply] `some`
145+
[apply] `Option.Rel.some`
142146
-/
143147
#guard_msgs in example : some true = (some true).map id := by
144148
change _ = .some _

tests/lean/run/coinductive_syntax.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -24,12 +24,12 @@ theorem casesOnTest (r : α → α → Prop) (a : α) : infSeq α r a → ∃ b,
2424
/--
2525
error: Invalid pattern: Expected a constructor or constant marked with `[match_pattern]`
2626
27-
Hint: These are similar:
28-
'Lean.Order.iterates.below.step',
29-
'Lean.Order.iterates.step',
30-
'Nat.le.below.step',
31-
'Nat.le.step',
32-
'infSeq._functor.step'
27+
Hint: Using one of these would be valid:
28+
[apply] `Nat.le.step`
29+
[apply] `Nat.le.below.step`
30+
[apply] `Lean.Order.iterates.below.step`
31+
[apply] `Lean.Order.iterates.step`
32+
[apply] `infSeq._functor.step`
3333
---
3434
error: Case tag `rhs` not found.
3535

0 commit comments

Comments
 (0)