1 /* Title: Pure/PIDE/markup.scala
5 Quasi-abstract markup elements.
17 def apply(elems: Set[String]): Elements = new Elements(elems)
18 def apply(elems: String*): Elements = apply(Set(elems: _*))
19 val empty: Elements = apply()
21 new Elements(Set.empty)
23 override def apply(elem: String): Boolean = true
24 override def toString: String = "Elements.full"
28 sealed class Elements private[Markup](private val rep: Set[String])
30 def apply(elem: String): Boolean = rep.contains(elem)
31 def + (elem: String): Elements = new Elements(rep + elem)
32 def ++ (elems: Elements): Elements = new Elements(rep ++ elems.rep)
33 override def toString: String = rep.mkString("Elements(", ",", ")")
40 val Name = new Properties.String(NAME)
43 val Kind = new Properties.String(KIND)
46 val Serial = new Properties.Long(SERIAL)
48 val INSTANCE = "instance"
49 val Instance = new Properties.String(INSTANCE)
54 val Empty = Markup("", Nil)
55 val Broken = Markup("broken", Nil)
57 class Markup_String(val name: String, prop: String)
59 private val Prop = new Properties.String(prop)
61 def apply(s: String): Markup = Markup(name, Prop(s))
62 def unapply(markup: Markup): Option[String] =
63 if (markup.name == name) Prop.unapply(markup.properties) else None
66 class Markup_Int(val name: String, prop: String)
68 private val Prop = new Properties.Int(prop)
70 def apply(i: Int): Markup = Markup(name, Prop(i))
71 def unapply(markup: Markup): Option[Int] =
72 if (markup.name == name) Prop.unapply(markup.properties) else None
75 class Markup_Long(val name: String, prop: String)
77 private val Prop = new Properties.Long(prop)
79 def apply(i: Long): Markup = Markup(name, Prop(i))
80 def unapply(markup: Markup): Option[Long] =
81 if (markup.name == name) Prop.unapply(markup.properties) else None
87 val BINDING = "binding"
92 val Def = new Properties.Long("def")
93 val Ref = new Properties.Long("ref")
95 def unapply(markup: Markup): Option[(String, String)] =
97 case Markup(ENTITY, props) =>
98 (props, props) match {
99 case (Kind(kind), Name(name)) => Some((kind, name))
109 val COMPLETION = "completion"
110 val NO_COMPLETION = "no_completion"
116 val END_LINE = "line"
117 val OFFSET = "offset"
118 val END_OFFSET = "end_offset"
122 val DEF_LINE = "def_line"
123 val DEF_OFFSET = "def_offset"
124 val DEF_END_OFFSET = "def_end_offset"
125 val DEF_FILE = "def_file"
126 val DEF_ID = "def_id"
128 val POSITION_PROPERTIES = Set(LINE, OFFSET, END_OFFSET, FILE, ID)
129 val POSITION = "position"
134 val EXPRESSION = "expression"
137 def unapply(markup: Markup): Option[String] =
139 case Markup(EXPRESSION, Kind(kind)) => Some(kind)
140 case Markup(EXPRESSION, _) => Some("")
148 val CITATION = "citation"
149 val Citation = new Markup_String(CITATION, NAME)
152 /* embedded languages */
154 val Symbols = new Properties.Boolean("symbols")
155 val Antiquotes = new Properties.Boolean("antiquotes")
156 val Delimited = new Properties.Boolean("delimited")
158 val LANGUAGE = "language"
164 val UNKNOWN = "unknown"
166 def unapply(markup: Markup): Option[(String, Boolean, Boolean, Boolean)] =
168 case Markup(LANGUAGE, props) =>
169 (props, props, props, props) match {
170 case (Name(name), Symbols(symbols), Antiquotes(antiquotes), Delimited(delimited)) =>
171 Some((name, symbols, antiquotes, delimited))
179 /* external resources */
182 val Path = new Markup_String(PATH, NAME)
185 val Url = new Markup_String(URL, NAME)
188 val Doc = new Markup_String(DOC, NAME)
191 /* pretty printing */
193 val Consistent = new Properties.Boolean("consistent")
194 val Indent = new Properties.Int("indent")
195 val Width = new Properties.Int("width")
200 def apply(c: Boolean, i: Int): Markup =
202 (if (c) Consistent(c) else Nil) :::
203 (if (i != 0) Indent(i) else Nil))
204 def unapply(markup: Markup): Option[(Boolean, Int)] =
205 if (markup.name == name) {
206 val c = Consistent.unapply(markup.properties).getOrElse(false)
207 val i = Indent.unapply(markup.properties).getOrElse(0)
216 def apply(w: Int, i: Int): Markup =
218 (if (w != 0) Width(w) else Nil) :::
219 (if (i != 0) Indent(i) else Nil))
220 def unapply(markup: Markup): Option[(Int, Int)] =
221 if (markup.name == name) {
222 val w = Width.unapply(markup.properties).getOrElse(0)
223 val i = Indent.unapply(markup.properties).getOrElse(0)
230 val BULLET = "bullet"
232 val SEPARATOR = "separator"
235 /* text properties */
239 val HIDDEN = "hidden"
245 val TYPE_NAME = "type_name"
248 val CONSTANT = "constant"
249 val DYNAMIC_FACT = "dynamic_fact"
257 val SKOLEM = "skolem"
260 val NUMERAL = "numeral"
261 val LITERAL = "literal"
262 val DELIMITER = "delimiter"
263 val INNER_STRING = "inner_string"
264 val INNER_CARTOUCHE = "inner_cartouche"
265 val INNER_COMMENT = "inner_comment"
267 val TOKEN_RANGE = "token_range"
269 val SORTING = "sorting"
270 val TYPING = "typing"
272 val ATTRIBUTE = "attribute"
273 val METHOD = "method"
278 val ANTIQUOTED = "antiquoted"
279 val ANTIQUOTE = "antiquote"
281 val ML_ANTIQUOTATION = "ML_antiquotation"
282 val DOCUMENT_ANTIQUOTATION = "document_antiquotation"
283 val DOCUMENT_ANTIQUOTATION_OPTION = "document_antiquotation_option"
288 val PARAGRAPH = "paragraph"
289 val TEXT_FOLD = "text_fold"
292 /* Markdown document structure */
294 val MARKDOWN_PARAGRAPH = "markdown_paragraph"
295 val Markdown_List = new Markup_String("markdown_list", "kind")
296 val Markdown_Item = new Markup_Int("markdown_item", "depth")
301 val ML_KEYWORD1 = "ML_keyword1"
302 val ML_KEYWORD2 = "ML_keyword2"
303 val ML_KEYWORD3 = "ML_keyword3"
304 val ML_DELIMITER = "ML_delimiter"
305 val ML_TVAR = "ML_tvar"
306 val ML_NUMERAL = "ML_numeral"
307 val ML_CHAR = "ML_char"
308 val ML_STRING = "ML_string"
309 val ML_COMMENT = "ML_comment"
310 val SML_STRING = "SML_string"
311 val SML_COMMENT = "SML_comment"
313 val ML_DEF = "ML_def"
314 val ML_OPEN = "ML_open"
315 val ML_STRUCTURE = "ML_structure"
316 val ML_TYPING = "ML_typing"
318 val ML_BREAKPOINT = "ML_breakpoint"
323 val COMMAND = "command"
324 val KEYWORD1 = "keyword1"
325 val KEYWORD2 = "keyword2"
326 val KEYWORD3 = "keyword3"
327 val QUASI_KEYWORD = "quasi_keyword"
328 val IMPROPER = "improper"
329 val OPERATOR = "operator"
330 val STRING = "string"
331 val ALT_STRING = "alt_string"
332 val VERBATIM = "verbatim"
333 val CARTOUCHE = "cartouche"
334 val COMMENT = "comment"
339 val Elapsed = new Properties.Double("elapsed")
340 val CPU = new Properties.Double("cpu")
341 val GC = new Properties.Double("gc")
343 object Timing_Properties
345 def apply(timing: isabelle.Timing): Properties.T =
346 Elapsed(timing.elapsed.seconds) ::: CPU(timing.cpu.seconds) ::: GC(timing.gc.seconds)
348 def unapply(props: Properties.T): Option[isabelle.Timing] =
349 (props, props, props) match {
350 case (Elapsed(elapsed), CPU(cpu), GC(gc)) =>
351 Some(new isabelle.Timing(Time.seconds(elapsed), Time.seconds(cpu), Time.seconds(gc)))
356 val TIMING = "timing"
360 def apply(timing: isabelle.Timing): Markup = Markup(TIMING, Timing_Properties(timing))
362 def unapply(markup: Markup): Option[isabelle.Timing] =
364 case Markup(TIMING, Timing_Properties(timing)) => Some(timing)
372 val COMMAND_TIMING = "command_timing"
377 val SUBGOALS = "subgoals"
378 val PROOF_STATE = "proof_state"
381 val SUBGOAL = "subgoal"
388 val ACCEPTED = "accepted"
389 val FORKED = "forked"
390 val JOINED = "joined"
391 val RUNNING = "running"
392 val FINISHED = "finished"
393 val FAILED = "failed"
396 /* interactive documents */
398 val VERSION = "version"
399 val ASSIGN = "assign"
404 val PROVER_COMMAND = "prover_command"
405 val PROVER_ARG = "prover_arg"
411 val STATUS = "status"
412 val REPORT = "report"
413 val RESULT = "result"
414 val WRITELN = "writeln"
416 val INFORMATION = "information"
417 val TRACING = "tracing"
418 val WARNING = "warning"
419 val LEGACY = "legacy"
421 val PROTOCOL = "protocol"
422 val SYSTEM = "system"
423 val STDOUT = "stdout"
424 val STDERR = "stderr"
427 val WRITELN_MESSAGE = "writeln_message"
428 val STATE_MESSAGE = "state_message"
429 val INFORMATION_MESSAGE = "information_message"
430 val TRACING_MESSAGE = "tracing_message"
431 val WARNING_MESSAGE = "warning_message"
432 val LEGACY_MESSAGE = "legacy_message"
433 val ERROR_MESSAGE = "error_message"
436 WRITELN -> WRITELN_MESSAGE,
437 STATE -> STATE_MESSAGE,
438 INFORMATION -> INFORMATION_MESSAGE,
439 TRACING -> TRACING_MESSAGE,
440 WARNING -> WARNING_MESSAGE,
441 LEGACY -> LEGACY_MESSAGE,
442 ERROR -> ERROR_MESSAGE)
444 val message: String => String = messages.withDefault((s: String) => s)
446 val Return_Code = new Properties.Int("return_code")
448 val NO_REPORT = "no_report"
452 val INTENSIFY = "intensify"
457 val BROWSER = "browser"
458 val GRAPHVIEW = "graphview"
460 val SENDBACK = "sendback"
461 val PADDING = "padding"
462 val PADDING_LINE = (PADDING, "line")
463 val PADDING_COMMAND = (PADDING, "command")
465 val DIALOG = "dialog"
466 val Result = new Properties.String(RESULT)
469 /* protocol message functions */
471 val FUNCTION = "function"
472 val Function = new Properties.String(FUNCTION)
474 val Assign_Update: Properties.T = List((FUNCTION, "assign_update"))
475 val Removed_Versions: Properties.T = List((FUNCTION, "removed_versions"))
477 object Protocol_Handler
479 def unapply(props: Properties.T): Option[(String)] =
481 case List((FUNCTION, "protocol_handler"), (NAME, name)) => Some(name)
486 val INVOKE_SCALA = "invoke_scala"
489 def unapply(props: Properties.T): Option[(String, String)] =
491 case List((FUNCTION, INVOKE_SCALA), (NAME, name), (ID, id)) => Some((name, id))
496 val CANCEL_SCALA = "cancel_scala"
499 def unapply(props: Properties.T): Option[String] =
501 case List((FUNCTION, CANCEL_SCALA), (ID, id)) => Some(id)
508 def unapply(props: Properties.T): Option[Properties.T] =
510 case (FUNCTION, "ML_statistics") :: stats => Some(stats)
515 object Task_Statistics
517 def unapply(props: Properties.T): Option[Properties.T] =
519 case (FUNCTION, "task_statistics") :: stats => Some(stats)
524 val LOADING_THEORY = "loading_theory"
525 object Loading_Theory
527 def unapply(props: Properties.T): Option[String] =
529 case List((FUNCTION, LOADING_THEORY), (NAME, name)) => Some(name)
534 val BUILD_THEORIES_RESULT = "build_theories_result"
535 object Build_Theories_Result
537 def unapply(props: Properties.T): Option[String] =
539 case List((FUNCTION, BUILD_THEORIES_RESULT), (ID, id)) => Some(id)
544 val PRINT_OPERATIONS = "print_operations"
547 /* debugger output */
549 val DEBUGGER_STATE = "debugger_state"
550 object Debugger_State
552 def unapply(props: Properties.T): Option[String] =
554 case List((FUNCTION, DEBUGGER_STATE), (NAME, name)) => Some(name)
559 val DEBUGGER_OUTPUT = "debugger_output"
560 object Debugger_Output
562 def unapply(props: Properties.T): Option[String] =
564 case List((FUNCTION, DEBUGGER_OUTPUT), (NAME, name)) => Some(name)
570 /* simplifier trace */
572 val SIMP_TRACE_PANEL = "simp_trace_panel"
574 val SIMP_TRACE_LOG = "simp_trace_log"
575 val SIMP_TRACE_STEP = "simp_trace_step"
576 val SIMP_TRACE_RECURSE = "simp_trace_recurse"
577 val SIMP_TRACE_HINT = "simp_trace_hint"
578 val SIMP_TRACE_IGNORE = "simp_trace_ignore"
580 val SIMP_TRACE_CANCEL = "simp_trace_cancel"
581 object Simp_Trace_Cancel
583 def unapply(props: Properties.T): Option[Long] =
585 case (FUNCTION, SIMP_TRACE_CANCEL) :: Serial(i) => Some(i)
592 sealed case class Markup(name: String, properties: Properties.T)
594 def markup(s: String): String =
595 YXML.string_of_tree(XML.Elem(this, List(XML.Text(s))))