Skip to content

Commit

Permalink
fix: Clear Diagnostics when file is closed (#1591)
Browse files Browse the repository at this point in the history
  • Loading branch information
lovettchris authored Oct 8, 2022
1 parent 4597422 commit 3eeb064
Show file tree
Hide file tree
Showing 7 changed files with 28 additions and 19 deletions.
17 changes: 17 additions & 0 deletions src/Lean/Data/Lsp/Ipc.lean
Original file line number Diff line number Diff line change
Expand Up @@ -39,6 +39,23 @@ def writeRequest (r : Request α) : IpcM Unit := do
def writeNotification (n : Notification α) : IpcM Unit := do
(←stdin).writeLspNotification n

def shutdown (requestNo : Nat) : IpcM Unit := do
let hIn ← stdout
let hOut ← stdin
hOut.writeLspRequest ⟨requestNo, "shutdown", Json.null⟩
repeat
let shutMsg ← hIn.readLspMessage
match shutMsg with
| Message.response id result =>
assert! result.isNull
if id != requestNo then
throw <| IO.userError s!"Expected id {requestNo}, got id {id}"

hOut.writeLspNotification ⟨"exit", Json.null⟩
break
| _ => -- ignore other messages in between.
pure ()

def readMessage : IpcM JsonRpc.Message := do
(←stdout).readLspMessage

Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Server/FileWorker.lean
Original file line number Diff line number Diff line change
Expand Up @@ -449,7 +449,7 @@ section MainLoop
handleRequest id method (toJson params)
mainLoop
| Message.notification "exit" none =>
let doc := (←get).doc
let doc := st.doc
doc.cancelTk.set
return ()
| Message.notification method (some params) =>
Expand Down
5 changes: 4 additions & 1 deletion src/Lean/Server/Watchdog.lean
Original file line number Diff line number Diff line change
Expand Up @@ -253,8 +253,11 @@ section ServerM
if exitCode = 0 then
-- Worker was terminated
fw.errorPendingRequests o ErrorCode.contentModified
("The file worker has been terminated. Either the header has changed,"
(s!"The file worker for {fw.doc.meta.uri} has been terminated. Either the header has changed,"
++ " or the file was closed, or the server is shutting down.")
-- one last message to clear the diagnostics for this file so that stale errors
-- do not remain in the editor forever.
publishDiagnostics fw.doc.meta #[] o
return WorkerEvent.terminated
else
-- Worker crashed
Expand Down
6 changes: 2 additions & 4 deletions tests/lean/interactive/run.lean
Original file line number Diff line number Diff line change
Expand Up @@ -178,8 +178,6 @@ partial def main (args : List String) : IO Unit := do
| _ =>
lastActualLineNo := lineNo
lineNo := lineNo + 1
Ipc.writeRequest ⟨requestNo, "shutdown", Json.null⟩
let shutResp ← Ipc.readResponseAs requestNo Json
assert! shutResp.result.isNull
Ipc.writeNotification ⟨"exit", Json.null⟩

Ipc.shutdown requestNo
discard <| Ipc.waitForExit
5 changes: 1 addition & 4 deletions tests/lean/server/diags.lean
Original file line number Diff line number Diff line change
Expand Up @@ -25,8 +25,5 @@ def main : IO Unit := do
else
throw $ userError "Failed parsing test file."

Ipc.writeRequest ⟨2, "shutdown", Json.null⟩
let shutResp ← Ipc.readResponseAs 2 Json
assert! shutResp.result.isNull
Ipc.writeNotification ⟨"exit", Json.null⟩
Ipc.shutdown 2
discard $ Ipc.waitForExit
5 changes: 1 addition & 4 deletions tests/lean/server/edits.lean
Original file line number Diff line number Diff line change
Expand Up @@ -30,8 +30,5 @@ def main : IO Unit := do
else
throw $ userError "Failed parsing test file."

Ipc.writeRequest ⟨3, "shutdown", Json.null⟩
let shutResp ← Ipc.readResponseAs 3 Json
assert! shutResp.result.isNull
Ipc.writeNotification ⟨"exit", Json.null⟩
Ipc.shutdown 3
discard $ Ipc.waitForExit
7 changes: 2 additions & 5 deletions tests/lean/server/init_exit.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,10 +8,7 @@ def main : IO Unit := do
hIn.flush
let initResp ← Ipc.readResponseAs 0 InitializeResult
let regWatchReq ← Ipc.readRequestAs "client/registerCapability" Json
Ipc.writeNotification ⟨"initialized", InitializedParams.mk⟩
Ipc.writeNotification ⟨"initialized", InitializedParams.mk⟩

Ipc.writeRequest ⟨1, "shutdown", Json.null⟩
let shutdownResp ← Ipc.readResponseAs 1 Json
assert! shutdownResp.result.isNull
Ipc.writeNotification ⟨"exit", Json.null⟩
Ipc.shutdown 1
discard Ipc.waitForExit

0 comments on commit 3eeb064

Please sign in to comment.