author | Thomas Lindae <thomas.lindae@in.tum.de> |
Wed, 24 Apr 2024 18:48:24 +0200 | |
changeset 81021 | 89bfada3a16d |
parent 75393 | 87ebf5a50283 |
child 81026 | c4bc259393f6 |
permissions | -rw-r--r-- |
72761 | 1 |
/* Title: Tools/VSCode/src/lsp.scala |
64605 | 2 |
Author: Makarius |
3 |
||
75201
8f6b8a46f54c
clarified modules: more uniform .scala vs. ts (amending 4519eeefe3b5);
wenzelm
parents:
75199
diff
changeset
|
4 |
Message formats for Language Server Protocol, with adhoc PIDE extensions. |
8f6b8a46f54c
clarified modules: more uniform .scala vs. ts (amending 4519eeefe3b5);
wenzelm
parents:
75199
diff
changeset
|
5 |
See https://github.com/Microsoft/language-server-protocol/blob/master/protocol.md |
64605 | 6 |
*/ |
7 |
||
8 |
package isabelle.vscode |
|
9 |
||
10 |
||
11 |
import isabelle._ |
|
12 |
||
64777
ca09695eb43c
clarified Document.Node.Name (again): canonical platform file;
wenzelm
parents:
64767
diff
changeset
|
13 |
import java.io.{File => JFile} |
ca09695eb43c
clarified Document.Node.Name (again): canonical platform file;
wenzelm
parents:
64767
diff
changeset
|
14 |
|
ca09695eb43c
clarified Document.Node.Name (again): canonical platform file;
wenzelm
parents:
64767
diff
changeset
|
15 |
|
75393 | 16 |
object LSP { |
64605 | 17 |
/* abstract message */ |
18 |
||
75393 | 19 |
object Message { |
67850 | 20 |
val empty: JSON.Object.T = JSON.Object("jsonrpc" -> "2.0") |
65922 | 21 |
|
75393 | 22 |
def log(prefix: String, json: JSON.T, logger: Logger, verbose: Boolean): Unit = { |
67851 | 23 |
val header = |
65922 | 24 |
json match { |
67851 | 25 |
case JSON.Object(obj) => obj -- (obj.keySet - "method" - "id") |
67850 | 26 |
case _ => JSON.Object.empty |
65922 | 27 |
} |
28 |
if (verbose || header.isEmpty) logger(prefix + "\n" + JSON.Format(json)) |
|
29 |
else logger(prefix + " " + JSON.Format(header)) |
|
30 |
} |
|
64605 | 31 |
} |
32 |
||
33 |
||
34 |
/* notification */ |
|
35 |
||
75393 | 36 |
object Notification { |
64605 | 37 |
def apply(method: String, params: JSON.T): JSON.T = |
38 |
Message.empty + ("method" -> method) + ("params" -> params) |
|
39 |
||
40 |
def unapply(json: JSON.T): Option[(String, Option[JSON.T])] = |
|
41 |
for { |
|
42 |
method <- JSON.string(json, "method") |
|
43 |
params = JSON.value(json, "params") |
|
44 |
} yield (method, params) |
|
45 |
} |
|
46 |
||
75393 | 47 |
class Notification0(name: String) { |
64605 | 48 |
def unapply(json: JSON.T): Option[Unit] = |
49 |
json match { |
|
50 |
case Notification(method, _) if method == name => Some(()) |
|
51 |
case _ => None |
|
52 |
} |
|
53 |
} |
|
54 |
||
55 |
||
56 |
/* request message */ |
|
57 |
||
75393 | 58 |
object Id { |
66675 | 59 |
def empty: Id = Id("") |
60 |
} |
|
61 |
||
75393 | 62 |
sealed case class Id(id: Any) { |
64605 | 63 |
require( |
64 |
id.isInstanceOf[Int] || |
|
65 |
id.isInstanceOf[Long] || |
|
66 |
id.isInstanceOf[Double] || |
|
73120
c3589f2dff31
more informative errors: simplify diagnosis of spurious failures reported by users;
wenzelm
parents:
72761
diff
changeset
|
67 |
id.isInstanceOf[String], "bad id type") |
64605 | 68 |
|
69 |
override def toString: String = id.toString |
|
70 |
} |
|
71 |
||
75393 | 72 |
object RequestMessage { |
64605 | 73 |
def apply(id: Id, method: String, params: JSON.T): JSON.T = |
74 |
Message.empty + ("id" -> id.id) + ("method" -> method) + ("params" -> params) |
|
75 |
||
76 |
def unapply(json: JSON.T): Option[(Id, String, Option[JSON.T])] = |
|
77 |
for { |
|
78 |
id <- JSON.long(json, "id") orElse JSON.double(json, "id") orElse JSON.string(json, "id") |
|
79 |
method <- JSON.string(json, "method") |
|
80 |
params = JSON.value(json, "params") |
|
81 |
} yield (Id(id), method, params) |
|
82 |
} |
|
83 |
||
75393 | 84 |
class Request0(name: String) { |
64605 | 85 |
def unapply(json: JSON.T): Option[Id] = |
86 |
json match { |
|
87 |
case RequestMessage(id, method, _) if method == name => Some(id) |
|
88 |
case _ => None |
|
89 |
} |
|
90 |
} |
|
91 |
||
75393 | 92 |
class RequestTextDocumentPosition(name: String) { |
64651 | 93 |
def unapply(json: JSON.T): Option[(Id, Line.Node_Position)] = |
64648
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
94 |
json match { |
64651 | 95 |
case RequestMessage(id, method, Some(TextDocumentPosition(node_pos))) if method == name => |
96 |
Some((id, node_pos)) |
|
64648
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
97 |
case _ => None |
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
98 |
} |
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
99 |
} |
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
100 |
|
64605 | 101 |
|
102 |
/* response message */ |
|
103 |
||
75393 | 104 |
object ResponseMessage { |
64605 | 105 |
def apply(id: Id, result: Option[JSON.T] = None, error: Option[ResponseError] = None): JSON.T = |
106 |
Message.empty + ("id" -> id.id) ++ |
|
64878 | 107 |
JSON.optional("result" -> result) ++ |
108 |
JSON.optional("error" -> error.map(_.json)) |
|
64605 | 109 |
|
110 |
def strict(id: Id, result: Option[JSON.T] = None, error: String = ""): JSON.T = |
|
111 |
if (error == "") apply(id, result = result) |
|
112 |
else apply(id, error = Some(ResponseError(code = ErrorCodes.serverErrorEnd, message = error))) |
|
66675 | 113 |
|
114 |
def is_empty(json: JSON.T): Boolean = |
|
115 |
JSON.string(json, "id") == Some("") && JSON.value(json, "result").isDefined |
|
64605 | 116 |
} |
117 |
||
75393 | 118 |
sealed case class ResponseError(code: Int, message: String, data: Option[JSON.T] = None) { |
64605 | 119 |
def json: JSON.T = |
67850 | 120 |
JSON.Object("code" -> code, "message" -> message) ++ JSON.optional("data" -> data) |
64605 | 121 |
} |
122 |
||
75393 | 123 |
object ErrorCodes { |
64605 | 124 |
val ParseError = -32700 |
125 |
val InvalidRequest = -32600 |
|
126 |
val MethodNotFound = -32601 |
|
127 |
val InvalidParams = -32602 |
|
128 |
val InternalError = -32603 |
|
129 |
val serverErrorStart = -32099 |
|
130 |
val serverErrorEnd = -32000 |
|
131 |
} |
|
132 |
||
133 |
||
134 |
/* init and exit */ |
|
135 |
||
75393 | 136 |
object Initialize extends Request0("initialize") { |
64605 | 137 |
def reply(id: Id, error: String): JSON.T = |
67850 | 138 |
ResponseMessage.strict( |
139 |
id, Some(JSON.Object("capabilities" -> ServerCapabilities.json)), error) |
|
64605 | 140 |
} |
141 |
||
75393 | 142 |
object ServerCapabilities { |
64605 | 143 |
val json: JSON.T = |
67850 | 144 |
JSON.Object( |
65160 | 145 |
"textDocumentSync" -> 2, |
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
146 |
"completionProvider" -> JSON.Object( |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
147 |
"resolveProvider" -> false, |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
148 |
"triggerCharacters" -> |
75199 | 149 |
Symbol.symbols.entries.flatMap(_.abbrevs).flatMap(_.toList).map(_.toString).distinct |
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
150 |
), |
64648
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
151 |
"hoverProvider" -> true, |
64767 | 152 |
"definitionProvider" -> true, |
153 |
"documentHighlightProvider" -> true) |
|
64605 | 154 |
} |
155 |
||
65229 | 156 |
object Initialized extends Notification0("initialized") |
157 |
||
75393 | 158 |
object Shutdown extends Request0("shutdown") { |
64605 | 159 |
def reply(id: Id, error: String): JSON.T = |
160 |
ResponseMessage.strict(id, Some("OK"), error) |
|
161 |
} |
|
162 |
||
163 |
object Exit extends Notification0("exit") |
|
164 |
||
165 |
||
166 |
/* document positions */ |
|
167 |
||
75393 | 168 |
object Position { |
64605 | 169 |
def apply(pos: Line.Position): JSON.T = |
67850 | 170 |
JSON.Object("line" -> pos.line, "character" -> pos.column) |
64605 | 171 |
|
172 |
def unapply(json: JSON.T): Option[Line.Position] = |
|
173 |
for { |
|
174 |
line <- JSON.int(json, "line") |
|
175 |
column <- JSON.int(json, "character") |
|
176 |
} yield Line.Position(line, column) |
|
177 |
} |
|
178 |
||
75393 | 179 |
object Range { |
65173 | 180 |
def compact(range: Line.Range): List[Int] = |
181 |
List(range.start.line, range.start.column, range.stop.line, range.stop.column) |
|
182 |
||
64605 | 183 |
def apply(range: Line.Range): JSON.T = |
67850 | 184 |
JSON.Object("start" -> Position(range.start), "end" -> Position(range.stop)) |
64605 | 185 |
|
186 |
def unapply(json: JSON.T): Option[Line.Range] = |
|
187 |
(JSON.value(json, "start"), JSON.value(json, "end")) match { |
|
188 |
case (Some(Position(start)), Some(Position(stop))) => Some(Line.Range(start, stop)) |
|
189 |
case _ => None |
|
190 |
} |
|
191 |
} |
|
192 |
||
75393 | 193 |
object Location { |
64651 | 194 |
def apply(loc: Line.Node_Range): JSON.T = |
67850 | 195 |
JSON.Object("uri" -> Url.print_file_name(loc.name), "range" -> Range(loc.range)) |
64605 | 196 |
|
64651 | 197 |
def unapply(json: JSON.T): Option[Line.Node_Range] = |
64605 | 198 |
for { |
199 |
uri <- JSON.string(json, "uri") |
|
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
200 |
if Url.is_wellformed_file(uri) |
64605 | 201 |
range_json <- JSON.value(json, "range") |
202 |
range <- Range.unapply(range_json) |
|
66235
d4fa51e7c4ff
retain symlinks in file names from VSCode: relevant for proper file locations in decorations etc.;
wenzelm
parents:
66216
diff
changeset
|
203 |
} yield Line.Node_Range(Url.absolute_file_name(uri), range) |
64605 | 204 |
} |
205 |
||
75393 | 206 |
object TextDocumentPosition { |
64651 | 207 |
def unapply(json: JSON.T): Option[Line.Node_Position] = |
64605 | 208 |
for { |
209 |
doc <- JSON.value(json, "textDocument") |
|
210 |
uri <- JSON.string(doc, "uri") |
|
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
211 |
if Url.is_wellformed_file(uri) |
64605 | 212 |
pos_json <- JSON.value(json, "position") |
213 |
pos <- Position.unapply(pos_json) |
|
66235
d4fa51e7c4ff
retain symlinks in file names from VSCode: relevant for proper file locations in decorations etc.;
wenzelm
parents:
66216
diff
changeset
|
214 |
} yield Line.Node_Position(Url.absolute_file_name(uri), pos) |
64605 | 215 |
} |
216 |
||
217 |
||
65103 | 218 |
/* marked strings */ |
65093 | 219 |
|
75393 | 220 |
sealed case class MarkedString(text: String, language: String = "plaintext") { |
67850 | 221 |
def json: JSON.T = JSON.Object("language" -> language, "value" -> text) |
65093 | 222 |
} |
223 |
||
75393 | 224 |
object MarkedStrings { |
65103 | 225 |
def json(msgs: List[MarkedString]): Option[JSON.T] = |
226 |
msgs match { |
|
227 |
case Nil => None |
|
228 |
case List(msg) => Some(msg.json) |
|
229 |
case _ => Some(msgs.map(_.json)) |
|
230 |
} |
|
231 |
} |
|
232 |
||
65093 | 233 |
|
64605 | 234 |
/* diagnostic messages */ |
235 |
||
75393 | 236 |
object MessageType { |
64605 | 237 |
val Error = 1 |
238 |
val Warning = 2 |
|
239 |
val Info = 3 |
|
240 |
val Log = 4 |
|
241 |
} |
|
242 |
||
75393 | 243 |
object DisplayMessage { |
73541 | 244 |
def apply(message_type: Int, message: String, show: Boolean): JSON.T = |
64605 | 245 |
Notification(if (show) "window/showMessage" else "window/logMessage", |
67850 | 246 |
JSON.Object("type" -> message_type, "message" -> message)) |
64605 | 247 |
} |
248 |
||
249 |
||
66138 | 250 |
/* commands */ |
251 |
||
75393 | 252 |
sealed case class Command(title: String, command: String, arguments: List[JSON.T] = Nil) { |
66138 | 253 |
def json: JSON.T = |
67850 | 254 |
JSON.Object("title" -> title, "command" -> command, "arguments" -> arguments) |
66138 | 255 |
} |
256 |
||
257 |
||
64605 | 258 |
/* document edits */ |
259 |
||
75393 | 260 |
object DidOpenTextDocument { |
64777
ca09695eb43c
clarified Document.Node.Name (again): canonical platform file;
wenzelm
parents:
64767
diff
changeset
|
261 |
def unapply(json: JSON.T): Option[(JFile, String, Long, String)] = |
64605 | 262 |
json match { |
263 |
case Notification("textDocument/didOpen", Some(params)) => |
|
264 |
for { |
|
265 |
doc <- JSON.value(params, "textDocument") |
|
266 |
uri <- JSON.string(doc, "uri") |
|
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
267 |
if Url.is_wellformed_file(uri) |
64605 | 268 |
lang <- JSON.string(doc, "languageId") |
269 |
version <- JSON.long(doc, "version") |
|
270 |
text <- JSON.string(doc, "text") |
|
66235
d4fa51e7c4ff
retain symlinks in file names from VSCode: relevant for proper file locations in decorations etc.;
wenzelm
parents:
66216
diff
changeset
|
271 |
} yield (Url.absolute_file(uri), lang, version, text) |
64605 | 272 |
case _ => None |
273 |
} |
|
274 |
} |
|
275 |
||
276 |
||
65160 | 277 |
sealed case class TextDocumentChange(range: Option[Line.Range], text: String) |
64605 | 278 |
|
75393 | 279 |
object DidChangeTextDocument { |
65160 | 280 |
def unapply_change(json: JSON.T): Option[TextDocumentChange] = |
281 |
for { text <- JSON.string(json, "text") } |
|
75204 | 282 |
yield TextDocumentChange(JSON.value(json, "range", Range.unapply), text) |
65160 | 283 |
|
284 |
def unapply(json: JSON.T): Option[(JFile, Long, List[TextDocumentChange])] = |
|
64605 | 285 |
json match { |
286 |
case Notification("textDocument/didChange", Some(params)) => |
|
287 |
for { |
|
288 |
doc <- JSON.value(params, "textDocument") |
|
289 |
uri <- JSON.string(doc, "uri") |
|
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
290 |
if Url.is_wellformed_file(uri) |
64605 | 291 |
version <- JSON.long(doc, "version") |
75204 | 292 |
changes <- JSON.list(params, "contentChanges", unapply_change) |
66235
d4fa51e7c4ff
retain symlinks in file names from VSCode: relevant for proper file locations in decorations etc.;
wenzelm
parents:
66216
diff
changeset
|
293 |
} yield (Url.absolute_file(uri), version, changes) |
64605 | 294 |
case _ => None |
295 |
} |
|
296 |
} |
|
297 |
||
75393 | 298 |
class TextDocumentNotification(name: String) { |
64777
ca09695eb43c
clarified Document.Node.Name (again): canonical platform file;
wenzelm
parents:
64767
diff
changeset
|
299 |
def unapply(json: JSON.T): Option[JFile] = |
64605 | 300 |
json match { |
301 |
case Notification(method, Some(params)) if method == name => |
|
302 |
for { |
|
303 |
doc <- JSON.value(params, "textDocument") |
|
304 |
uri <- JSON.string(doc, "uri") |
|
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
305 |
if Url.is_wellformed_file(uri) |
66235
d4fa51e7c4ff
retain symlinks in file names from VSCode: relevant for proper file locations in decorations etc.;
wenzelm
parents:
66216
diff
changeset
|
306 |
} yield Url.absolute_file(uri) |
64605 | 307 |
case _ => None |
308 |
} |
|
309 |
} |
|
310 |
||
311 |
object DidCloseTextDocument extends TextDocumentNotification("textDocument/didClose") |
|
312 |
object DidSaveTextDocument extends TextDocumentNotification("textDocument/didSave") |
|
313 |
||
314 |
||
66675 | 315 |
/* workspace edits */ |
316 |
||
75393 | 317 |
sealed case class TextEdit(range: Line.Range, new_text: String) { |
67850 | 318 |
def json: JSON.T = JSON.Object("range" -> Range(range), "newText" -> new_text) |
66675 | 319 |
} |
320 |
||
75393 | 321 |
sealed case class TextDocumentEdit(file: JFile, version: Long, edits: List[TextEdit]) { |
66675 | 322 |
def json: JSON.T = |
67850 | 323 |
JSON.Object( |
324 |
"textDocument" -> JSON.Object("uri" -> Url.print_file(file), "version" -> version), |
|
66675 | 325 |
"edits" -> edits.map(_.json)) |
326 |
} |
|
327 |
||
75393 | 328 |
object WorkspaceEdit { |
66675 | 329 |
def apply(edits: List[TextDocumentEdit]): JSON.T = |
330 |
RequestMessage(Id.empty, "workspace/applyEdit", |
|
67850 | 331 |
JSON.Object("edit" -> JSON.Object("documentChanges" -> edits.map(_.json)))) |
66675 | 332 |
} |
333 |
||
334 |
||
64877 | 335 |
/* completion */ |
336 |
||
337 |
sealed case class CompletionItem( |
|
338 |
label: String, |
|
339 |
kind: Option[Int] = None, |
|
340 |
detail: Option[String] = None, |
|
341 |
documentation: Option[String] = None, |
|
66140 | 342 |
text: Option[String] = None, |
66138 | 343 |
range: Option[Line.Range] = None, |
75393 | 344 |
command: Option[Command] = None |
345 |
) { |
|
64877 | 346 |
def json: JSON.T = |
67850 | 347 |
JSON.Object("label" -> label) ++ |
66061 | 348 |
JSON.optional("kind" -> kind) ++ |
349 |
JSON.optional("detail" -> detail) ++ |
|
350 |
JSON.optional("documentation" -> documentation) ++ |
|
66140 | 351 |
JSON.optional("insertText" -> text) ++ |
66062
175772371cd0
use old-style "textEdit" for the sake of the external protocol (see also vscode-languageserver-node/issues/188);
wenzelm
parents:
66061
diff
changeset
|
352 |
JSON.optional("range" -> range.map(Range(_))) ++ |
66675 | 353 |
JSON.optional("textEdit" -> range.map(r => TextEdit(r, text.getOrElse(label)).json)) ++ |
66138 | 354 |
JSON.optional("command" -> command.map(_.json)) |
64877 | 355 |
} |
356 |
||
75393 | 357 |
object Completion extends RequestTextDocumentPosition("textDocument/completion") { |
64877 | 358 |
def reply(id: Id, result: List[CompletionItem]): JSON.T = |
359 |
ResponseMessage(id, Some(result.map(_.json))) |
|
360 |
} |
|
361 |
||
362 |
||
66138 | 363 |
/* spell checker */ |
364 |
||
75393 | 365 |
object Include_Word extends Notification0("PIDE/include_word") { |
66139
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
366 |
val command = Command("Include word", "isabelle.include-word") |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
367 |
} |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
368 |
|
75393 | 369 |
object Include_Word_Permanently extends Notification0("PIDE/include_word_permanently") { |
66139
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
370 |
val command = Command("Include word permanently", "isabelle.include-word-permanently") |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
371 |
} |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
372 |
|
75393 | 373 |
object Exclude_Word extends Notification0("PIDE/exclude_word") { |
66139
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
374 |
val command = Command("Exclude word", "isabelle.exclude-word") |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
375 |
} |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
376 |
|
75393 | 377 |
object Exclude_Word_Permanently extends Notification0("PIDE/exclude_word_permanently") { |
66139
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
378 |
val command = Command("Exclude word permanently", "isabelle.exclude-word-permanently") |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
379 |
} |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
380 |
|
75393 | 381 |
object Reset_Words extends Notification0("PIDE/reset_words") { |
66139
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
382 |
val command = Command("Reset non-permanent words", "isabelle.reset-words") |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
383 |
} |
66138 | 384 |
|
385 |
||
64605 | 386 |
/* hover request */ |
387 |
||
75393 | 388 |
object Hover extends RequestTextDocumentPosition("textDocument/hover") { |
389 |
def reply(id: Id, result: Option[(Line.Range, List[MarkedString])]): JSON.T = { |
|
64605 | 390 |
val res = |
391 |
result match { |
|
64747 | 392 |
case Some((range, contents)) => |
67850 | 393 |
JSON.Object( |
394 |
"contents" -> MarkedStrings.json(contents).getOrElse(Nil), |
|
395 |
"range" -> Range(range)) |
|
396 |
case None => JSON.Object("contents" -> Nil) |
|
64605 | 397 |
} |
398 |
ResponseMessage(id, Some(res)) |
|
399 |
} |
|
400 |
} |
|
64648
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
401 |
|
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
402 |
|
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
403 |
/* goto definition request */ |
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
404 |
|
75393 | 405 |
object GotoDefinition extends RequestTextDocumentPosition("textDocument/definition") { |
64651 | 406 |
def reply(id: Id, result: List[Line.Node_Range]): JSON.T = |
75204 | 407 |
ResponseMessage(id, Some(result.map(Location.apply))) |
64648
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
408 |
} |
64675 | 409 |
|
410 |
||
64767 | 411 |
/* document highlights request */ |
412 |
||
75393 | 413 |
object DocumentHighlight { |
64767 | 414 |
def text(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(1)) |
415 |
def read(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(2)) |
|
416 |
def write(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(3)) |
|
417 |
} |
|
418 |
||
75393 | 419 |
sealed case class DocumentHighlight(range: Line.Range, kind: Option[Int] = None) { |
64767 | 420 |
def json: JSON.T = |
421 |
kind match { |
|
67850 | 422 |
case None => JSON.Object("range" -> Range(range)) |
423 |
case Some(k) => JSON.Object("range" -> Range(range), "kind" -> k) |
|
64767 | 424 |
} |
425 |
} |
|
426 |
||
75393 | 427 |
object DocumentHighlights extends RequestTextDocumentPosition("textDocument/documentHighlight") { |
64767 | 428 |
def reply(id: Id, result: List[DocumentHighlight]): JSON.T = |
429 |
ResponseMessage(id, Some(result.map(_.json))) |
|
430 |
} |
|
431 |
||
432 |
||
64675 | 433 |
/* diagnostics */ |
434 |
||
75393 | 435 |
sealed case class Diagnostic( |
436 |
range: Line.Range, |
|
437 |
message: String, |
|
438 |
severity: Option[Int] = None, |
|
439 |
code: Option[Int] = None, |
|
440 |
source: Option[String] = None |
|
441 |
) { |
|
64675 | 442 |
def json: JSON.T = |
443 |
Message.empty + ("range" -> Range(range)) + ("message" -> message) ++ |
|
64878 | 444 |
JSON.optional("severity" -> severity) ++ |
445 |
JSON.optional("code" -> code) ++ |
|
446 |
JSON.optional("source" -> source) |
|
64675 | 447 |
} |
448 |
||
75393 | 449 |
object DiagnosticSeverity { |
64762
cd545bec3fd0
clarified message severity, based on empirical observation of VSCode 1.8.1;
wenzelm
parents:
64747
diff
changeset
|
450 |
val Error = 1 |
cd545bec3fd0
clarified message severity, based on empirical observation of VSCode 1.8.1;
wenzelm
parents:
64747
diff
changeset
|
451 |
val Warning = 2 |
cd545bec3fd0
clarified message severity, based on empirical observation of VSCode 1.8.1;
wenzelm
parents:
64747
diff
changeset
|
452 |
val Information = 3 |
cd545bec3fd0
clarified message severity, based on empirical observation of VSCode 1.8.1;
wenzelm
parents:
64747
diff
changeset
|
453 |
val Hint = 4 |
64675 | 454 |
} |
455 |
||
75393 | 456 |
object PublishDiagnostics { |
64777
ca09695eb43c
clarified Document.Node.Name (again): canonical platform file;
wenzelm
parents:
64767
diff
changeset
|
457 |
def apply(file: JFile, diagnostics: List[Diagnostic]): JSON.T = |
64675 | 458 |
Notification("textDocument/publishDiagnostics", |
67850 | 459 |
JSON.Object("uri" -> Url.print_file(file), "diagnostics" -> diagnostics.map(_.json))) |
64675 | 460 |
} |
65094 | 461 |
|
462 |
||
463 |
/* decorations */ |
|
464 |
||
75393 | 465 |
sealed case class Decoration_Options(range: Line.Range, hover_message: List[MarkedString]) { |
65094 | 466 |
def json: JSON.T = |
67850 | 467 |
JSON.Object("range" -> Range.compact(range)) ++ |
65173 | 468 |
JSON.optional("hover_message" -> MarkedStrings.json(hover_message)) |
65094 | 469 |
} |
470 |
||
75393 | 471 |
sealed case class Decoration(decorations: List[(String, List[Decoration_Options])]) { |
65095 | 472 |
def json(file: JFile): JSON.T = |
65094 | 473 |
Notification("PIDE/decoration", |
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
474 |
JSON.Object( |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
475 |
"uri" -> Url.print_file(file), |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
476 |
"entries" -> decorations.map(decoration => JSON.Object( |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
477 |
"type" -> decoration._1, |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
478 |
"content" -> decoration._2.map(_.json)) |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
479 |
)) |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
480 |
) |
65094 | 481 |
} |
65189 | 482 |
|
483 |
||
66215
9a8b6b86350c
proper hyperlink_command, notably for locate_query;
wenzelm
parents:
66211
diff
changeset
|
484 |
/* caret update: bidirectional */ |
65189 | 485 |
|
75393 | 486 |
object Caret_Update { |
66216 | 487 |
def apply(node_pos: Line.Node_Position, focus: Boolean): JSON.T = |
66215
9a8b6b86350c
proper hyperlink_command, notably for locate_query;
wenzelm
parents:
66211
diff
changeset
|
488 |
Notification("PIDE/caret_update", |
67850 | 489 |
JSON.Object( |
490 |
"uri" -> Url.print_file_name(node_pos.name), |
|
66215
9a8b6b86350c
proper hyperlink_command, notably for locate_query;
wenzelm
parents:
66211
diff
changeset
|
491 |
"line" -> node_pos.pos.line, |
66216 | 492 |
"character" -> node_pos.pos.column, |
493 |
"focus" -> focus)) |
|
66215
9a8b6b86350c
proper hyperlink_command, notably for locate_query;
wenzelm
parents:
66211
diff
changeset
|
494 |
|
65189 | 495 |
def unapply(json: JSON.T): Option[Option[(JFile, Line.Position)]] = |
496 |
json match { |
|
497 |
case Notification("PIDE/caret_update", Some(params)) => |
|
498 |
val caret = |
|
499 |
for { |
|
500 |
uri <- JSON.string(params, "uri") |
|
501 |
if Url.is_wellformed_file(uri) |
|
502 |
pos <- Position.unapply(params) |
|
66235
d4fa51e7c4ff
retain symlinks in file names from VSCode: relevant for proper file locations in decorations etc.;
wenzelm
parents:
66216
diff
changeset
|
503 |
} yield (Url.absolute_file(uri), pos) |
65189 | 504 |
Some(caret) |
505 |
case _ => None |
|
506 |
} |
|
507 |
} |
|
508 |
||
65201 | 509 |
|
510 |
/* dynamic output */ |
|
511 |
||
75393 | 512 |
object Dynamic_Output { |
66090 | 513 |
def apply(content: String): JSON.T = |
67850 | 514 |
Notification("PIDE/dynamic_output", JSON.Object("content" -> content)) |
65189 | 515 |
} |
65977
c51b74be23b6
provide preview content on Scala side (similar to output);
wenzelm
parents:
65924
diff
changeset
|
516 |
|
c51b74be23b6
provide preview content on Scala side (similar to output);
wenzelm
parents:
65924
diff
changeset
|
517 |
|
66096 | 518 |
/* state output */ |
519 |
||
75393 | 520 |
object State_Output { |
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
521 |
def apply(id: Counter.ID, content: String, auto_update: Boolean): JSON.T = |
75393 | 522 |
Notification("PIDE/state_output", |
523 |
JSON.Object("id" -> id, "content" -> content, "auto_update" -> auto_update)) |
|
66096 | 524 |
} |
525 |
||
81021
89bfada3a16d
lsp: added State_Init_Response message;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
75393
diff
changeset
|
526 |
object State_Init_Response { |
89bfada3a16d
lsp: added State_Init_Response message;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
75393
diff
changeset
|
527 |
def apply(id: Counter.ID): JSON.T = |
89bfada3a16d
lsp: added State_Init_Response message;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
75393
diff
changeset
|
528 |
Notification("PIDE/state_init_response", JSON.Object("id" -> id)) |
89bfada3a16d
lsp: added State_Init_Response message;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
75393
diff
changeset
|
529 |
} |
89bfada3a16d
lsp: added State_Init_Response message;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
75393
diff
changeset
|
530 |
|
75393 | 531 |
class State_Id_Notification(name: String) { |
66096 | 532 |
def unapply(json: JSON.T): Option[Counter.ID] = |
533 |
json match { |
|
534 |
case Notification(method, Some(params)) if method == name => JSON.long(params, "id") |
|
535 |
case _ => None |
|
536 |
} |
|
537 |
} |
|
538 |
||
539 |
object State_Init extends Notification0("PIDE/state_init") |
|
540 |
object State_Exit extends State_Id_Notification("PIDE/state_exit") |
|
541 |
object State_Locate extends State_Id_Notification("PIDE/state_locate") |
|
542 |
object State_Update extends State_Id_Notification("PIDE/state_update") |
|
543 |
||
75393 | 544 |
object State_Auto_Update { |
66211 | 545 |
def unapply(json: JSON.T): Option[(Counter.ID, Boolean)] = |
546 |
json match { |
|
547 |
case Notification("PIDE/state_auto_update", Some(params)) => |
|
548 |
for { |
|
549 |
id <- JSON.long(params, "id") |
|
550 |
enabled <- JSON.bool(params, "enabled") |
|
551 |
} yield (id, enabled) |
|
552 |
case _ => None |
|
553 |
} |
|
554 |
} |
|
555 |
||
66096 | 556 |
|
65983 | 557 |
/* preview */ |
65977
c51b74be23b6
provide preview content on Scala side (similar to output);
wenzelm
parents:
65924
diff
changeset
|
558 |
|
75393 | 559 |
object Preview_Request { |
65983 | 560 |
def unapply(json: JSON.T): Option[(JFile, Int)] = |
561 |
json match { |
|
562 |
case Notification("PIDE/preview_request", Some(params)) => |
|
563 |
for { |
|
564 |
uri <- JSON.string(params, "uri") |
|
565 |
if Url.is_wellformed_file(uri) |
|
566 |
column <- JSON.int(params, "column") |
|
66235
d4fa51e7c4ff
retain symlinks in file names from VSCode: relevant for proper file locations in decorations etc.;
wenzelm
parents:
66216
diff
changeset
|
567 |
} yield (Url.absolute_file(uri), column) |
65983 | 568 |
case _ => None |
569 |
} |
|
570 |
} |
|
571 |
||
75393 | 572 |
object Preview_Response { |
65983 | 573 |
def apply(file: JFile, column: Int, label: String, content: String): JSON.T = |
574 |
Notification("PIDE/preview_response", |
|
67850 | 575 |
JSON.Object( |
65983 | 576 |
"uri" -> Url.print_file(file), |
577 |
"column" -> column, |
|
578 |
"label" -> label, |
|
579 |
"content" -> content)) |
|
65977
c51b74be23b6
provide preview content on Scala side (similar to output);
wenzelm
parents:
65924
diff
changeset
|
580 |
} |
64605 | 581 |
} |