Skip to content

Commit

Permalink
fix: don't issue atomic id completions when there is a dangling dot
Browse files Browse the repository at this point in the history
  • Loading branch information
mhuisi committed Oct 25, 2024
1 parent b5e276f commit 51aa9e7
Show file tree
Hide file tree
Showing 3 changed files with 32 additions and 10 deletions.
24 changes: 14 additions & 10 deletions src/Lean/Server/Completion/CompletionCollectors.lean
Original file line number Diff line number Diff line change
Expand Up @@ -182,7 +182,9 @@ end Utils

section IdCompletionUtils

private def matchAtomic (id : Name) (declName : Name) : Option Float :=
private def matchAtomic (id : Name) (declName : Name) (danglingDot : Bool) : Option Float := do
if danglingDot then
none
match id, declName with
| .str .anonymous s₁, .str .anonymous s₂ => fuzzyMatchScoreWithThreshold? s₁ s₂
| _, _ => none
Expand Down Expand Up @@ -360,7 +362,7 @@ private def idCompletionCore
if id.isAtomic then
-- search for matches in the local context
for localDecl in (← getLCtx) do
if let some score := matchAtomic id localDecl.userName then
if let some score := matchAtomic id localDecl.userName danglingDot then
addUnresolvedCompletionItem localDecl.userName (.fvar localDecl.fvarId) (kind := CompletionItemKind.variable) score
-- search for matches in the environment
let env ← getEnv
Expand Down Expand Up @@ -395,10 +397,11 @@ private def idCompletionCore
matchUsingNamespace Name.anonymous
if let some (bestLabel, bestScore) := bestMatch? then
addUnresolvedCompletionItem bestLabel (.const declName) (← getCompletionKindForDecl c) bestScore
-- Recall that aliases may not be atomic and include the namespace where they were created.
let matchAlias (ns : Name) (alias : Name) : Option Float :=
-- Recall that aliases may not be atomic and include the namespace where they were created.
if ns.isPrefixOf alias then
matchAtomic id (alias.replacePrefix ns Name.anonymous)
let alias := alias.replacePrefix ns Name.anonymous
matchAtomic id alias danglingDot
else
none
let eligibleHeaderDecls ← getEligibleHeaderDecls env
Expand All @@ -412,7 +415,7 @@ private def idCompletionCore
match openDecl with
| OpenDecl.explicit openedId resolvedId =>
if allowCompletion eligibleHeaderDecls env resolvedId then
if let some score := matchAtomic id openedId then
if let some score := matchAtomic id openedId danglingDot then
addUnresolvedCompletionItemForDecl (.mkSimple openedId.getString!) resolvedId score
| OpenDecl.simple ns _ =>
getAliasState env |>.forM fun alias declNames => do
Expand All @@ -430,11 +433,12 @@ private def idCompletionCore
| _ => return ()
searchAlias ctx.currNamespace
-- Search keywords
if let .str .anonymous s := id then
let keywords := Parser.getTokenTable env
for keyword in keywords.findPrefix s do
if let some score := fuzzyMatchScoreWithThreshold? s keyword then
addKeywordCompletionItem keyword score
if !danglingDot then
if let .str .anonymous s := id then
let keywords := Parser.getTokenTable env
for keyword in keywords.findPrefix s do
if let some score := fuzzyMatchScoreWithThreshold? s keyword then
addKeywordCompletionItem keyword score
-- Search namespaces
completeNamespaces ctx id danglingDot

Expand Down
12 changes: 12 additions & 0 deletions tests/lean/interactive/completionDanglingDot.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,12 @@
-- Regression test for a bug where the dangling dot was not accounted for in some
-- atomic completions, which lead to invalid completions after a dangling dot

def foo : Unit :=
x. -- No completions expected
--^ textDocument/completion

def bar : Array Nat := Id.run do
let mut x := sorry
let foo := x. -- No completions expected
--^ textDocument/completion
sorry
Original file line number Diff line number Diff line change
@@ -0,0 +1,6 @@
{"textDocument": {"uri": "file:///completionDanglingDot.lean"},
"position": {"line": 4, "character": 4}}
{"items": [], "isIncomplete": true}
{"textDocument": {"uri": "file:///completionDanglingDot.lean"},
"position": {"line": 9, "character": 15}}
{"items": [], "isIncomplete": true}

0 comments on commit 51aa9e7

Please sign in to comment.