author | Thomas Lindae <thomas.lindae@in.tum.de> |
Thu, 16 May 2024 11:59:33 +0200 | |
changeset 81042 | f1e0ca5aaa6b |
parent 81039 | 9c5a4ba4ced6 |
child 81043 | 2174ec5f0d0c |
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) |
|
81026
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
112 |
else apply(id, error = Some(ResponseError(code = ErrorCodes.jsonrpcReservedErrorRangeEnd, 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 |
|
81026
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
129 |
|
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
130 |
val jsonrpcReservedErrorRangeStart = -32099 |
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
131 |
val ServerNotInitialized = -32002 |
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
132 |
val UnknownErrorCode = -32001 |
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
133 |
val jsonrpcReservedErrorRangeEnd = -32000 |
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
134 |
|
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
135 |
val lspReservedErrorRangeStart = -32899 |
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
136 |
val RequestFailed = -32803 |
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
137 |
val ServerCancelled = -32802 |
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
138 |
val ContentModified = -32801 |
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
139 |
val RequestCancelled = -32800 |
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
140 |
val lspReservedErrorRangeEnd = -32800 |
64605 | 141 |
} |
81026
c4bc259393f6
lsp: updated ErrorCodes;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81021
diff
changeset
|
142 |
|
64605 | 143 |
|
144 |
/* init and exit */ |
|
145 |
||
75393 | 146 |
object Initialize extends Request0("initialize") { |
64605 | 147 |
def reply(id: Id, error: String): JSON.T = |
67850 | 148 |
ResponseMessage.strict( |
149 |
id, Some(JSON.Object("capabilities" -> ServerCapabilities.json)), error) |
|
64605 | 150 |
} |
151 |
||
75393 | 152 |
object ServerCapabilities { |
64605 | 153 |
val json: JSON.T = |
67850 | 154 |
JSON.Object( |
65160 | 155 |
"textDocumentSync" -> 2, |
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
156 |
"completionProvider" -> JSON.Object( |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
157 |
"resolveProvider" -> false, |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
158 |
"triggerCharacters" -> |
75199 | 159 |
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
|
160 |
), |
64648
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
161 |
"hoverProvider" -> true, |
64767 | 162 |
"definitionProvider" -> true, |
163 |
"documentHighlightProvider" -> true) |
|
64605 | 164 |
} |
165 |
||
65229 | 166 |
object Initialized extends Notification0("initialized") |
167 |
||
75393 | 168 |
object Shutdown extends Request0("shutdown") { |
64605 | 169 |
def reply(id: Id, error: String): JSON.T = |
170 |
ResponseMessage.strict(id, Some("OK"), error) |
|
171 |
} |
|
172 |
||
173 |
object Exit extends Notification0("exit") |
|
174 |
||
175 |
||
176 |
/* document positions */ |
|
177 |
||
75393 | 178 |
object Position { |
64605 | 179 |
def apply(pos: Line.Position): JSON.T = |
67850 | 180 |
JSON.Object("line" -> pos.line, "character" -> pos.column) |
64605 | 181 |
|
182 |
def unapply(json: JSON.T): Option[Line.Position] = |
|
183 |
for { |
|
184 |
line <- JSON.int(json, "line") |
|
185 |
column <- JSON.int(json, "character") |
|
186 |
} yield Line.Position(line, column) |
|
187 |
} |
|
188 |
||
75393 | 189 |
object Range { |
65173 | 190 |
def compact(range: Line.Range): List[Int] = |
191 |
List(range.start.line, range.start.column, range.stop.line, range.stop.column) |
|
192 |
||
64605 | 193 |
def apply(range: Line.Range): JSON.T = |
67850 | 194 |
JSON.Object("start" -> Position(range.start), "end" -> Position(range.stop)) |
64605 | 195 |
|
196 |
def unapply(json: JSON.T): Option[Line.Range] = |
|
197 |
(JSON.value(json, "start"), JSON.value(json, "end")) match { |
|
198 |
case (Some(Position(start)), Some(Position(stop))) => Some(Line.Range(start, stop)) |
|
199 |
case _ => None |
|
200 |
} |
|
201 |
} |
|
202 |
||
75393 | 203 |
object Location { |
64651 | 204 |
def apply(loc: Line.Node_Range): JSON.T = |
67850 | 205 |
JSON.Object("uri" -> Url.print_file_name(loc.name), "range" -> Range(loc.range)) |
64605 | 206 |
|
64651 | 207 |
def unapply(json: JSON.T): Option[Line.Node_Range] = |
64605 | 208 |
for { |
209 |
uri <- JSON.string(json, "uri") |
|
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
210 |
if Url.is_wellformed_file(uri) |
64605 | 211 |
range_json <- JSON.value(json, "range") |
212 |
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
|
213 |
} yield Line.Node_Range(Url.absolute_file_name(uri), range) |
64605 | 214 |
} |
215 |
||
75393 | 216 |
object TextDocumentPosition { |
64651 | 217 |
def unapply(json: JSON.T): Option[Line.Node_Position] = |
64605 | 218 |
for { |
219 |
doc <- JSON.value(json, "textDocument") |
|
220 |
uri <- JSON.string(doc, "uri") |
|
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
221 |
if Url.is_wellformed_file(uri) |
64605 | 222 |
pos_json <- JSON.value(json, "position") |
223 |
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
|
224 |
} yield Line.Node_Position(Url.absolute_file_name(uri), pos) |
64605 | 225 |
} |
226 |
||
227 |
||
65103 | 228 |
/* marked strings */ |
65093 | 229 |
|
75393 | 230 |
sealed case class MarkedString(text: String, language: String = "plaintext") { |
67850 | 231 |
def json: JSON.T = JSON.Object("language" -> language, "value" -> text) |
65093 | 232 |
} |
233 |
||
75393 | 234 |
object MarkedStrings { |
65103 | 235 |
def json(msgs: List[MarkedString]): Option[JSON.T] = |
236 |
msgs match { |
|
237 |
case Nil => None |
|
238 |
case List(msg) => Some(msg.json) |
|
239 |
case _ => Some(msgs.map(_.json)) |
|
240 |
} |
|
241 |
} |
|
242 |
||
65093 | 243 |
|
64605 | 244 |
/* diagnostic messages */ |
245 |
||
75393 | 246 |
object MessageType { |
64605 | 247 |
val Error = 1 |
248 |
val Warning = 2 |
|
249 |
val Info = 3 |
|
250 |
val Log = 4 |
|
251 |
} |
|
252 |
||
75393 | 253 |
object DisplayMessage { |
73541 | 254 |
def apply(message_type: Int, message: String, show: Boolean): JSON.T = |
64605 | 255 |
Notification(if (show) "window/showMessage" else "window/logMessage", |
67850 | 256 |
JSON.Object("type" -> message_type, "message" -> message)) |
64605 | 257 |
} |
258 |
||
259 |
||
66138 | 260 |
/* commands */ |
261 |
||
75393 | 262 |
sealed case class Command(title: String, command: String, arguments: List[JSON.T] = Nil) { |
66138 | 263 |
def json: JSON.T = |
67850 | 264 |
JSON.Object("title" -> title, "command" -> command, "arguments" -> arguments) |
66138 | 265 |
} |
266 |
||
267 |
||
64605 | 268 |
/* document edits */ |
269 |
||
75393 | 270 |
object DidOpenTextDocument { |
64777
ca09695eb43c
clarified Document.Node.Name (again): canonical platform file;
wenzelm
parents:
64767
diff
changeset
|
271 |
def unapply(json: JSON.T): Option[(JFile, String, Long, String)] = |
64605 | 272 |
json match { |
273 |
case Notification("textDocument/didOpen", Some(params)) => |
|
274 |
for { |
|
275 |
doc <- JSON.value(params, "textDocument") |
|
276 |
uri <- JSON.string(doc, "uri") |
|
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
277 |
if Url.is_wellformed_file(uri) |
64605 | 278 |
lang <- JSON.string(doc, "languageId") |
279 |
version <- JSON.long(doc, "version") |
|
280 |
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
|
281 |
} yield (Url.absolute_file(uri), lang, version, text) |
64605 | 282 |
case _ => None |
283 |
} |
|
284 |
} |
|
285 |
||
286 |
||
65160 | 287 |
sealed case class TextDocumentChange(range: Option[Line.Range], text: String) |
64605 | 288 |
|
75393 | 289 |
object DidChangeTextDocument { |
65160 | 290 |
def unapply_change(json: JSON.T): Option[TextDocumentChange] = |
291 |
for { text <- JSON.string(json, "text") } |
|
75204 | 292 |
yield TextDocumentChange(JSON.value(json, "range", Range.unapply), text) |
65160 | 293 |
|
294 |
def unapply(json: JSON.T): Option[(JFile, Long, List[TextDocumentChange])] = |
|
64605 | 295 |
json match { |
296 |
case Notification("textDocument/didChange", Some(params)) => |
|
297 |
for { |
|
298 |
doc <- JSON.value(params, "textDocument") |
|
299 |
uri <- JSON.string(doc, "uri") |
|
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
300 |
if Url.is_wellformed_file(uri) |
64605 | 301 |
version <- JSON.long(doc, "version") |
75204 | 302 |
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
|
303 |
} yield (Url.absolute_file(uri), version, changes) |
64605 | 304 |
case _ => None |
305 |
} |
|
306 |
} |
|
307 |
||
75393 | 308 |
class TextDocumentNotification(name: String) { |
64777
ca09695eb43c
clarified Document.Node.Name (again): canonical platform file;
wenzelm
parents:
64767
diff
changeset
|
309 |
def unapply(json: JSON.T): Option[JFile] = |
64605 | 310 |
json match { |
311 |
case Notification(method, Some(params)) if method == name => |
|
312 |
for { |
|
313 |
doc <- JSON.value(params, "textDocument") |
|
314 |
uri <- JSON.string(doc, "uri") |
|
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
315 |
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
|
316 |
} yield Url.absolute_file(uri) |
64605 | 317 |
case _ => None |
318 |
} |
|
319 |
} |
|
320 |
||
321 |
object DidCloseTextDocument extends TextDocumentNotification("textDocument/didClose") |
|
322 |
object DidSaveTextDocument extends TextDocumentNotification("textDocument/didSave") |
|
323 |
||
324 |
||
66675 | 325 |
/* workspace edits */ |
326 |
||
75393 | 327 |
sealed case class TextEdit(range: Line.Range, new_text: String) { |
67850 | 328 |
def json: JSON.T = JSON.Object("range" -> Range(range), "newText" -> new_text) |
66675 | 329 |
} |
330 |
||
75393 | 331 |
sealed case class TextDocumentEdit(file: JFile, version: Long, edits: List[TextEdit]) { |
66675 | 332 |
def json: JSON.T = |
67850 | 333 |
JSON.Object( |
334 |
"textDocument" -> JSON.Object("uri" -> Url.print_file(file), "version" -> version), |
|
66675 | 335 |
"edits" -> edits.map(_.json)) |
336 |
} |
|
337 |
||
75393 | 338 |
object WorkspaceEdit { |
66675 | 339 |
def apply(edits: List[TextDocumentEdit]): JSON.T = |
340 |
RequestMessage(Id.empty, "workspace/applyEdit", |
|
67850 | 341 |
JSON.Object("edit" -> JSON.Object("documentChanges" -> edits.map(_.json)))) |
66675 | 342 |
} |
343 |
||
344 |
||
64877 | 345 |
/* completion */ |
346 |
||
347 |
sealed case class CompletionItem( |
|
348 |
label: String, |
|
349 |
kind: Option[Int] = None, |
|
350 |
detail: Option[String] = None, |
|
351 |
documentation: Option[String] = None, |
|
66140 | 352 |
text: Option[String] = None, |
66138 | 353 |
range: Option[Line.Range] = None, |
75393 | 354 |
command: Option[Command] = None |
355 |
) { |
|
64877 | 356 |
def json: JSON.T = |
67850 | 357 |
JSON.Object("label" -> label) ++ |
66061 | 358 |
JSON.optional("kind" -> kind) ++ |
359 |
JSON.optional("detail" -> detail) ++ |
|
360 |
JSON.optional("documentation" -> documentation) ++ |
|
66140 | 361 |
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
|
362 |
JSON.optional("range" -> range.map(Range(_))) ++ |
66675 | 363 |
JSON.optional("textEdit" -> range.map(r => TextEdit(r, text.getOrElse(label)).json)) ++ |
66138 | 364 |
JSON.optional("command" -> command.map(_.json)) |
64877 | 365 |
} |
366 |
||
75393 | 367 |
object Completion extends RequestTextDocumentPosition("textDocument/completion") { |
64877 | 368 |
def reply(id: Id, result: List[CompletionItem]): JSON.T = |
369 |
ResponseMessage(id, Some(result.map(_.json))) |
|
370 |
} |
|
371 |
||
372 |
||
66138 | 373 |
/* spell checker */ |
374 |
||
75393 | 375 |
object Include_Word extends Notification0("PIDE/include_word") { |
66139
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
376 |
val command = Command("Include word", "isabelle.include-word") |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
377 |
} |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
378 |
|
75393 | 379 |
object Include_Word_Permanently extends Notification0("PIDE/include_word_permanently") { |
66139
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
380 |
val command = Command("Include word permanently", "isabelle.include-word-permanently") |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
381 |
} |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
382 |
|
75393 | 383 |
object Exclude_Word extends Notification0("PIDE/exclude_word") { |
66139
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
384 |
val command = Command("Exclude word", "isabelle.exclude-word") |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
385 |
} |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
386 |
|
75393 | 387 |
object Exclude_Word_Permanently extends Notification0("PIDE/exclude_word_permanently") { |
66139
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
388 |
val command = Command("Exclude word permanently", "isabelle.exclude-word-permanently") |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
389 |
} |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
390 |
|
75393 | 391 |
object Reset_Words extends Notification0("PIDE/reset_words") { |
66139
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
392 |
val command = Command("Reset non-permanent words", "isabelle.reset-words") |
6a8f8be2741c
provide spell-checker menu via completion commands;
wenzelm
parents:
66138
diff
changeset
|
393 |
} |
66138 | 394 |
|
395 |
||
64605 | 396 |
/* hover request */ |
397 |
||
75393 | 398 |
object Hover extends RequestTextDocumentPosition("textDocument/hover") { |
399 |
def reply(id: Id, result: Option[(Line.Range, List[MarkedString])]): JSON.T = { |
|
64605 | 400 |
val res = |
401 |
result match { |
|
64747 | 402 |
case Some((range, contents)) => |
67850 | 403 |
JSON.Object( |
404 |
"contents" -> MarkedStrings.json(contents).getOrElse(Nil), |
|
405 |
"range" -> Range(range)) |
|
406 |
case None => JSON.Object("contents" -> Nil) |
|
64605 | 407 |
} |
408 |
ResponseMessage(id, Some(res)) |
|
409 |
} |
|
410 |
} |
|
64648
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
411 |
|
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
412 |
|
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
413 |
/* goto definition request */ |
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
414 |
|
75393 | 415 |
object GotoDefinition extends RequestTextDocumentPosition("textDocument/definition") { |
64651 | 416 |
def reply(id: Id, result: List[Line.Node_Range]): JSON.T = |
75204 | 417 |
ResponseMessage(id, Some(result.map(Location.apply))) |
64648
5d7f741aaccb
basic support for hyperlinks / Goto Definition Request;
wenzelm
parents:
64605
diff
changeset
|
418 |
} |
64675 | 419 |
|
420 |
||
64767 | 421 |
/* document highlights request */ |
422 |
||
75393 | 423 |
object DocumentHighlight { |
64767 | 424 |
def text(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(1)) |
425 |
def read(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(2)) |
|
426 |
def write(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(3)) |
|
427 |
} |
|
428 |
||
75393 | 429 |
sealed case class DocumentHighlight(range: Line.Range, kind: Option[Int] = None) { |
64767 | 430 |
def json: JSON.T = |
431 |
kind match { |
|
67850 | 432 |
case None => JSON.Object("range" -> Range(range)) |
433 |
case Some(k) => JSON.Object("range" -> Range(range), "kind" -> k) |
|
64767 | 434 |
} |
435 |
} |
|
436 |
||
75393 | 437 |
object DocumentHighlights extends RequestTextDocumentPosition("textDocument/documentHighlight") { |
64767 | 438 |
def reply(id: Id, result: List[DocumentHighlight]): JSON.T = |
439 |
ResponseMessage(id, Some(result.map(_.json))) |
|
440 |
} |
|
441 |
||
442 |
||
64675 | 443 |
/* diagnostics */ |
444 |
||
75393 | 445 |
sealed case class Diagnostic( |
446 |
range: Line.Range, |
|
447 |
message: String, |
|
448 |
severity: Option[Int] = None, |
|
449 |
code: Option[Int] = None, |
|
450 |
source: Option[String] = None |
|
451 |
) { |
|
64675 | 452 |
def json: JSON.T = |
453 |
Message.empty + ("range" -> Range(range)) + ("message" -> message) ++ |
|
64878 | 454 |
JSON.optional("severity" -> severity) ++ |
455 |
JSON.optional("code" -> code) ++ |
|
456 |
JSON.optional("source" -> source) |
|
64675 | 457 |
} |
458 |
||
75393 | 459 |
object DiagnosticSeverity { |
64762
cd545bec3fd0
clarified message severity, based on empirical observation of VSCode 1.8.1;
wenzelm
parents:
64747
diff
changeset
|
460 |
val Error = 1 |
cd545bec3fd0
clarified message severity, based on empirical observation of VSCode 1.8.1;
wenzelm
parents:
64747
diff
changeset
|
461 |
val Warning = 2 |
cd545bec3fd0
clarified message severity, based on empirical observation of VSCode 1.8.1;
wenzelm
parents:
64747
diff
changeset
|
462 |
val Information = 3 |
cd545bec3fd0
clarified message severity, based on empirical observation of VSCode 1.8.1;
wenzelm
parents:
64747
diff
changeset
|
463 |
val Hint = 4 |
64675 | 464 |
} |
465 |
||
75393 | 466 |
object PublishDiagnostics { |
64777
ca09695eb43c
clarified Document.Node.Name (again): canonical platform file;
wenzelm
parents:
64767
diff
changeset
|
467 |
def apply(file: JFile, diagnostics: List[Diagnostic]): JSON.T = |
64675 | 468 |
Notification("textDocument/publishDiagnostics", |
67850 | 469 |
JSON.Object("uri" -> Url.print_file(file), "diagnostics" -> diagnostics.map(_.json))) |
64675 | 470 |
} |
65094 | 471 |
|
472 |
||
473 |
/* decorations */ |
|
474 |
||
75393 | 475 |
sealed case class Decoration_Options(range: Line.Range, hover_message: List[MarkedString]) { |
65094 | 476 |
def json: JSON.T = |
67850 | 477 |
JSON.Object("range" -> Range.compact(range)) ++ |
65173 | 478 |
JSON.optional("hover_message" -> MarkedStrings.json(hover_message)) |
65094 | 479 |
} |
480 |
||
75393 | 481 |
sealed case class Decoration(decorations: List[(String, List[Decoration_Options])]) { |
65095 | 482 |
def json(file: JFile): JSON.T = |
65094 | 483 |
Notification("PIDE/decoration", |
75126
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
484 |
JSON.Object( |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
485 |
"uri" -> Url.print_file(file), |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
486 |
"entries" -> decorations.map(decoration => JSON.Object( |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
487 |
"type" -> decoration._1, |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
488 |
"content" -> decoration._2.map(_.json)) |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
489 |
)) |
da1108a6d249
various improvements to Isabelle/VSCode (by Denis Paluca and Fabian Huch);
wenzelm
parents:
73541
diff
changeset
|
490 |
) |
65094 | 491 |
} |
65189 | 492 |
|
493 |
||
66215
9a8b6b86350c
proper hyperlink_command, notably for locate_query;
wenzelm
parents:
66211
diff
changeset
|
494 |
/* caret update: bidirectional */ |
65189 | 495 |
|
75393 | 496 |
object Caret_Update { |
66216 | 497 |
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
|
498 |
Notification("PIDE/caret_update", |
67850 | 499 |
JSON.Object( |
500 |
"uri" -> Url.print_file_name(node_pos.name), |
|
66215
9a8b6b86350c
proper hyperlink_command, notably for locate_query;
wenzelm
parents:
66211
diff
changeset
|
501 |
"line" -> node_pos.pos.line, |
66216 | 502 |
"character" -> node_pos.pos.column, |
503 |
"focus" -> focus)) |
|
66215
9a8b6b86350c
proper hyperlink_command, notably for locate_query;
wenzelm
parents:
66211
diff
changeset
|
504 |
|
65189 | 505 |
def unapply(json: JSON.T): Option[Option[(JFile, Line.Position)]] = |
506 |
json match { |
|
507 |
case Notification("PIDE/caret_update", Some(params)) => |
|
508 |
val caret = |
|
509 |
for { |
|
510 |
uri <- JSON.string(params, "uri") |
|
511 |
if Url.is_wellformed_file(uri) |
|
512 |
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
|
513 |
} yield (Url.absolute_file(uri), pos) |
65189 | 514 |
Some(caret) |
515 |
case _ => None |
|
516 |
} |
|
517 |
} |
|
518 |
||
65201 | 519 |
|
520 |
/* dynamic output */ |
|
521 |
||
75393 | 522 |
object Dynamic_Output { |
81039
9c5a4ba4ced6
lsp: unified PIDE/decorations and dynamic output decorations format;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81037
diff
changeset
|
523 |
def apply(content: String, decorations: Option[List[(String, List[Decoration_Options])]] = None): JSON.T = |
81037
f1aef7110329
lsp: added decorations to dynamic output;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81032
diff
changeset
|
524 |
Notification("PIDE/dynamic_output", |
f1aef7110329
lsp: added decorations to dynamic output;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81032
diff
changeset
|
525 |
JSON.Object("content" -> content) ++ |
f1aef7110329
lsp: added decorations to dynamic output;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81032
diff
changeset
|
526 |
JSON.optional( |
f1aef7110329
lsp: added decorations to dynamic output;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81032
diff
changeset
|
527 |
"decorations" -> decorations.map(decorations => |
f1aef7110329
lsp: added decorations to dynamic output;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81032
diff
changeset
|
528 |
decorations.map(decoration => JSON.Object( |
81039
9c5a4ba4ced6
lsp: unified PIDE/decorations and dynamic output decorations format;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81037
diff
changeset
|
529 |
"type" -> decoration._1, |
9c5a4ba4ced6
lsp: unified PIDE/decorations and dynamic output decorations format;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81037
diff
changeset
|
530 |
"content" -> decoration._2.map(_.json)) |
81037
f1aef7110329
lsp: added decorations to dynamic output;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81032
diff
changeset
|
531 |
)) |
f1aef7110329
lsp: added decorations to dynamic output;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81032
diff
changeset
|
532 |
)) |
65189 | 533 |
} |
65977
c51b74be23b6
provide preview content on Scala side (similar to output);
wenzelm
parents:
65924
diff
changeset
|
534 |
|
81029
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
535 |
object Output_Set_Margin { |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
536 |
def unapply(json: JSON.T): Option[Double] = |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
537 |
json match { |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
538 |
case Notification("PIDE/output_set_margin", Some(params)) => |
81032 | 539 |
JSON.double(params, "margin") |
81029
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
540 |
case _ => None |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
541 |
} |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
542 |
} |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
543 |
|
65977
c51b74be23b6
provide preview content on Scala side (similar to output);
wenzelm
parents:
65924
diff
changeset
|
544 |
|
66096 | 545 |
/* state output */ |
546 |
||
75393 | 547 |
object State_Output { |
81042
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
548 |
def apply( |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
549 |
id: Counter.ID, |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
550 |
content: String, |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
551 |
auto_update: Boolean, |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
552 |
decorations: Option[List[(String, List[Decoration_Options])]] = None |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
553 |
): JSON.T = |
75393 | 554 |
Notification("PIDE/state_output", |
81042
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
555 |
JSON.Object("id" -> id, "content" -> content, "auto_update" -> auto_update) ++ |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
556 |
JSON.optional( |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
557 |
"decorations" -> decorations.map(decorations => |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
558 |
decorations.map(decoration => JSON.Object( |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
559 |
"type" -> decoration._1, |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
560 |
"content" -> decoration._2.map(_.json)) |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
561 |
)) |
f1e0ca5aaa6b
lsp: added decorations to state updates;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81039
diff
changeset
|
562 |
)) |
66096 | 563 |
} |
564 |
||
75393 | 565 |
class State_Id_Notification(name: String) { |
66096 | 566 |
def unapply(json: JSON.T): Option[Counter.ID] = |
567 |
json match { |
|
568 |
case Notification(method, Some(params)) if method == name => JSON.long(params, "id") |
|
569 |
case _ => None |
|
570 |
} |
|
571 |
} |
|
572 |
||
81028
84f6f17274d0
lsp: changed State_Init notification into a request and removed State_Init_Response;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81026
diff
changeset
|
573 |
object State_Init extends Request0("PIDE/state_init") { |
84f6f17274d0
lsp: changed State_Init notification into a request and removed State_Init_Response;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81026
diff
changeset
|
574 |
def reply(id: Id, state_id: Counter.ID): JSON.T = |
84f6f17274d0
lsp: changed State_Init notification into a request and removed State_Init_Response;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81026
diff
changeset
|
575 |
ResponseMessage(id, Some(JSON.Object("state_id" -> state_id))) |
84f6f17274d0
lsp: changed State_Init notification into a request and removed State_Init_Response;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81026
diff
changeset
|
576 |
} |
84f6f17274d0
lsp: changed State_Init notification into a request and removed State_Init_Response;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81026
diff
changeset
|
577 |
|
66096 | 578 |
object State_Exit extends State_Id_Notification("PIDE/state_exit") |
579 |
object State_Locate extends State_Id_Notification("PIDE/state_locate") |
|
580 |
object State_Update extends State_Id_Notification("PIDE/state_update") |
|
581 |
||
75393 | 582 |
object State_Auto_Update { |
66211 | 583 |
def unapply(json: JSON.T): Option[(Counter.ID, Boolean)] = |
584 |
json match { |
|
585 |
case Notification("PIDE/state_auto_update", Some(params)) => |
|
586 |
for { |
|
587 |
id <- JSON.long(params, "id") |
|
588 |
enabled <- JSON.bool(params, "enabled") |
|
589 |
} yield (id, enabled) |
|
590 |
case _ => None |
|
591 |
} |
|
592 |
} |
|
593 |
||
81029
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
594 |
object State_Set_Margin { |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
595 |
def unapply(json: JSON.T): Option[(Counter.ID, Double)] = |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
596 |
json match { |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
597 |
case Notification("PIDE/state_set_margin", Some(params)) => |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
598 |
for { |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
599 |
id <- JSON.long(params, "id") |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
600 |
margin <- JSON.double(params, "margin") |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
601 |
} yield (id, margin) |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
602 |
case _ => None |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
603 |
} |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
604 |
} |
f4cb1e35c63e
lsp: added Output_Set_Margin and State_Set_Margin Notifications;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81028
diff
changeset
|
605 |
|
66096 | 606 |
|
81030
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
607 |
/* symbols */ |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
608 |
|
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
609 |
object Symbols_Request extends Request0("PIDE/symbols_request") { |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
610 |
def reply(id: Id, symbols: Symbol.Symbols): JSON.T = { |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
611 |
def json(symbol: Symbol.Entry): JSON.T = |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
612 |
JSON.Object( |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
613 |
"symbol" -> symbol.symbol, |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
614 |
"name" -> symbol.name, |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
615 |
"argument" -> symbol.argument.toString, |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
616 |
) ++ |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
617 |
JSON.optional("code", symbol.code) ++ |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
618 |
JSON.optional("font", symbol.font) ++ |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
619 |
JSON.Object( |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
620 |
"groups" -> symbol.groups, |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
621 |
"abbrevs" -> symbol.abbrevs, |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
622 |
) |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
623 |
|
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
624 |
ResponseMessage(id, Some(symbols.entries.map(s => json(s)))) |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
625 |
} |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
626 |
|
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
627 |
} |
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
628 |
|
88879ff1cef5
lsp: added Symbols_Request;
Thomas Lindae <thomas.lindae@in.tum.de>
parents:
81029
diff
changeset
|
629 |
|
65983 | 630 |
/* preview */ |
65977
c51b74be23b6
provide preview content on Scala side (similar to output);
wenzelm
parents:
65924
diff
changeset
|
631 |
|
75393 | 632 |
object Preview_Request { |
65983 | 633 |
def unapply(json: JSON.T): Option[(JFile, Int)] = |
634 |
json match { |
|
635 |
case Notification("PIDE/preview_request", Some(params)) => |
|
636 |
for { |
|
637 |
uri <- JSON.string(params, "uri") |
|
638 |
if Url.is_wellformed_file(uri) |
|
639 |
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
|
640 |
} yield (Url.absolute_file(uri), column) |
65983 | 641 |
case _ => None |
642 |
} |
|
643 |
} |
|
644 |
||
75393 | 645 |
object Preview_Response { |
65983 | 646 |
def apply(file: JFile, column: Int, label: String, content: String): JSON.T = |
647 |
Notification("PIDE/preview_response", |
|
67850 | 648 |
JSON.Object( |
65983 | 649 |
"uri" -> Url.print_file(file), |
650 |
"column" -> column, |
|
651 |
"label" -> label, |
|
652 |
"content" -> content)) |
|
65977
c51b74be23b6
provide preview content on Scala side (similar to output);
wenzelm
parents:
65924
diff
changeset
|
653 |
} |
64605 | 654 |
} |