Skip to content

Commit 2594a8e

Browse files
authored
fix: namespace completion to only use the short name (#8350)
This PR changes namespace completion to use the same algorithm as declaration identifier completion, which makes it use the short name (last name component) for completions instead of the full name, avoiding namespace duplications. Closes #5654
1 parent b24e232 commit 2594a8e

File tree

3 files changed

+37
-64
lines changed

3 files changed

+37
-64
lines changed

src/Lean/Server/Completion/CompletionCollectors.lean

Lines changed: 34 additions & 61 deletions
Original file line numberDiff line numberDiff line change
@@ -209,43 +209,43 @@ section IdCompletionUtils
209209
(Name.mkStr p (s.extract 0 ⟨newLen - optDot - len⟩), newLen)
210210
(go id).1
211211

212-
def matchNamespace (ns : Name) (nsFragment : Name) (danglingDot : Bool) : Bool :=
213-
if danglingDot then
214-
if nsFragment != ns && nsFragment.isPrefixOf ns then
215-
true
216-
else
217-
false
218-
else
219-
match ns, nsFragment with
220-
| .str p₁ s₁, .str p₂ s₂ =>
221-
if p₁ == p₂ then s₂.charactersIn s₁ else false
222-
| _, _ => false
212+
def bestLabelForDecl? (ctx : ContextInfo) (declName : Name) (id : Name) (danglingDot : Bool) :
213+
M (Option Name) := Prod.snd <$> StateT.run (s := none) do
214+
let matchUsingNamespace (ns : Name) : StateT (Option Name) M Unit := do
215+
let some label ← matchDecl? ns id danglingDot declName
216+
| return
217+
modify fun
218+
| none =>
219+
some label
220+
| some bestLabel =>
221+
-- for open namespaces `A` and `A.B` and a decl `A.B.c`, pick the decl `c` over `B.c`
222+
if label.isSuffixOf bestLabel then
223+
some label
224+
else
225+
some bestLabel
226+
let rec visitNamespaces (ns : Name) : StateT (Option Name) M Unit := do
227+
let Name.str p .. := ns
228+
| return ()
229+
matchUsingNamespace ns
230+
visitNamespaces p
231+
-- use current namespace
232+
visitNamespaces ctx.currNamespace
233+
-- use open decls
234+
for openDecl in ctx.openDecls do
235+
let OpenDecl.simple ns exs := openDecl
236+
| pure ()
237+
if exs.contains declName then
238+
continue
239+
matchUsingNamespace ns
240+
matchUsingNamespace Name.anonymous
223241

224242
def completeNamespaces (ctx : ContextInfo) (id : Name) (danglingDot : Bool) : M Unit := do
225243
let env ← getEnv
226-
let add (ns : Name) (ns' : Name) : M Unit :=
227-
if danglingDot then
228-
addNamespaceCompletionItem (ns.replacePrefix (ns' ++ id) Name.anonymous)
229-
else
230-
addNamespaceCompletionItem (ns.replacePrefix ns' Name.anonymous)
231244
env.getNamespaceSet |>.forM fun ns => do
232245
unless ns.isInternal || env.contains ns do -- Ignore internal and namespaces that are also declaration names
233-
for openDecl in ctx.openDecls do
234-
match openDecl with
235-
| OpenDecl.simple ns' _ =>
236-
if matchNamespace ns (ns' ++ id) danglingDot then
237-
add ns ns'
238-
return ()
239-
| _ => pure ()
240-
-- use current namespace
241-
let rec visitNamespaces (ns' : Name) : M Unit := do
242-
if matchNamespace ns (ns' ++ id) danglingDot then
243-
add ns ns'
244-
else
245-
match ns' with
246-
| Name.str p .. => visitNamespaces p
247-
| _ => return ()
248-
visitNamespaces ctx.currNamespace
246+
let label? ← bestLabelForDecl? ctx ns id danglingDot
247+
if let some bestLabel := label? then
248+
addNamespaceCompletionItem bestLabel
249249

250250
end IdCompletionUtils
251251

@@ -258,7 +258,7 @@ section DotCompletionUtils
258258
| _ => false
259259
if isConstOf then
260260
return true
261-
let some e ← unfoldeDefinitionGuarded? e | return false
261+
let some e ← unfoldDefinitionGuarded? e | return false
262262
isDefEqToAppOf e declName
263263

264264
private def isDotCompletionMethod (typeName : Name) (info : ConstantInfo) : MetaM Bool :=
@@ -356,34 +356,7 @@ private def idCompletionCore
356356
-- search for matches in the environment
357357
let env ← getEnv
358358
forEligibleDeclsWithCancellationM fun declName c => do
359-
let bestMatch? ← (·.2) <$> StateT.run (s := none) do
360-
let matchUsingNamespace (ns : Name) : StateT (Option Name) M Unit := do
361-
let some label ← matchDecl? ns id danglingDot declName
362-
| return
363-
modify fun
364-
| none =>
365-
some label
366-
| some bestLabel =>
367-
-- for open namespaces `A` and `A.B` and a decl `A.B.c`, pick the decl `c` over `B.c`
368-
if label.isSuffixOf bestLabel then
369-
some label
370-
else
371-
some bestLabel
372-
let rec visitNamespaces (ns : Name) : StateT (Option Name) M Unit := do
373-
let Name.str p .. := ns
374-
| return ()
375-
matchUsingNamespace ns
376-
visitNamespaces p
377-
-- use current namespace
378-
visitNamespaces ctx.currNamespace
379-
-- use open decls
380-
for openDecl in ctx.openDecls do
381-
let OpenDecl.simple ns exs := openDecl
382-
| pure ()
383-
if exs.contains declName then
384-
continue
385-
matchUsingNamespace ns
386-
matchUsingNamespace Name.anonymous
359+
let bestMatch? ← bestLabelForDecl? ctx declName id danglingDot
387360
if let some bestLabel := bestMatch? then
388361
addUnresolvedCompletionItem bestLabel (.const declName) (← getCompletionKindForDecl c)
389362
RequestCancellation.check

src/Lean/Server/Completion/CompletionUtils.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -66,7 +66,7 @@ where
6666
else
6767
shortenIn id contextNamespace.getPrefix
6868

69-
def unfoldeDefinitionGuarded? (e : Expr) : MetaM (Option Expr) :=
69+
def unfoldDefinitionGuarded? (e : Expr) : MetaM (Option Expr) :=
7070
try Lean.Meta.unfoldDefinition? e catch _ => pure none
7171

7272
partial def getDotCompletionTypeNames (type : Expr) : MetaM (Array Name) :=
@@ -78,7 +78,7 @@ where
7878
if isStructure (← getEnv) typeName then
7979
for parentName in (← getAllParentStructures typeName) do
8080
modify fun s => s.push parentName
81-
let some type ← unfoldeDefinitionGuarded? type | return ()
81+
let some type ← unfoldDefinitionGuarded? type | return ()
8282
visit type
8383

8484
end Lean.Server.Completion

tests/lean/interactive/compNamespace.lean.expected.out

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -37,7 +37,7 @@
3737
{"textDocument": {"uri": "file:///compNamespace.lean"},
3838
"position": {"line": 16, "character": 16}}
3939
{"items":
40-
[{"label": "Foo.LongNamespaceExample",
40+
[{"label": "LongNamespaceExample",
4141
"kind": 9,
4242
"detail": "namespace",
4343
"data":

0 commit comments

Comments
 (0)