PIDE markup for Markdown items (which may consist of multiple paragraphs or lists);
1 /* Title: Pure/PIDE/markup.scala
4 Quasi-abstract markup elements.
16 def apply(elems: Set[String]): Elements = new Elements(elems)
17 def apply(elems: String*): Elements = apply(Set(elems: _*))
18 val empty: Elements = apply()
20 new Elements(Set.empty)
22 override def apply(elem: String): Boolean = true
23 override def toString: String = "Elements.full"
27 sealed class Elements private[Markup](private val rep: Set[String])
29 def apply(elem: String): Boolean = rep.contains(elem)
30 def + (elem: String): Elements = new Elements(rep + elem)
31 def ++ (elems: Elements): Elements = new Elements(rep ++ elems.rep)
32 def - (elem: String): Elements = new Elements(rep - elem)
33 def -- (elems: Elements): Elements = new Elements(rep -- elems.rep)
34 override def toString: String = rep.mkString("Elements(", ",", ")")
41 val Name = new Properties.String(NAME)
44 val Kind = new Properties.String(KIND)
46 val CONTENT = "content"
47 val Content = new Properties.String(CONTENT)
50 val Serial = new Properties.Long(SERIAL)
52 val INSTANCE = "instance"
53 val Instance = new Properties.String(INSTANCE)
58 val Empty = Markup("", Nil)
59 val Broken = Markup("broken", Nil)
61 class Markup_String(val name: String, prop: String)
63 private val Prop = new Properties.String(prop)
65 def apply(s: String): Markup = Markup(name, Prop(s))
66 def unapply(markup: Markup): Option[String] =
67 if (markup.name == name) Prop.unapply(markup.properties) else None
70 class Markup_Int(val name: String, prop: String)
72 private val Prop = new Properties.Int(prop)
74 def apply(i: Int): Markup = Markup(name, Prop(i))
75 def unapply(markup: Markup): Option[Int] =
76 if (markup.name == name) Prop.unapply(markup.properties) else None
79 class Markup_Long(val name: String, prop: String)
81 private val Prop = new Properties.Long(prop)
83 def apply(i: Long): Markup = Markup(name, Prop(i))
84 def unapply(markup: Markup): Option[Long] =
85 if (markup.name == name) Prop.unapply(markup.properties) else None
91 val BINDING = "binding"
96 val Def = new Properties.Long("def")
97 val Ref = new Properties.Long("ref")
99 def unapply(markup: Markup): Option[(String, String)] =
101 case Markup(ENTITY, props) =>
102 val kind = Kind.unapply(props).getOrElse("")
103 val name = Name.unapply(props).getOrElse("")
112 val COMPLETION = "completion"
113 val NO_COMPLETION = "no_completion"
119 val END_LINE = "line"
120 val OFFSET = "offset"
121 val END_OFFSET = "end_offset"
125 val DEF_LINE = "def_line"
126 val DEF_OFFSET = "def_offset"
127 val DEF_END_OFFSET = "def_end_offset"
128 val DEF_FILE = "def_file"
129 val DEF_ID = "def_id"
131 val POSITION_PROPERTIES = Set(LINE, OFFSET, END_OFFSET, FILE, ID)
132 val POSITION = "position"
137 val EXPRESSION = "expression"
140 def unapply(markup: Markup): Option[String] =
142 case Markup(EXPRESSION, Kind(kind)) => Some(kind)
143 case Markup(EXPRESSION, _) => Some("")
151 val CITATION = "citation"
152 val Citation = new Markup_String(CITATION, NAME)
155 /* embedded languages */
157 val Symbols = new Properties.Boolean("symbols")
158 val Antiquotes = new Properties.Boolean("antiquotes")
159 val Delimited = new Properties.Boolean("delimited")
161 val LANGUAGE = "language"
167 val UNKNOWN = "unknown"
169 def unapply(markup: Markup): Option[(String, Boolean, Boolean, Boolean)] =
171 case Markup(LANGUAGE, props) =>
172 (props, props, props, props) match {
173 case (Name(name), Symbols(symbols), Antiquotes(antiquotes), Delimited(delimited)) =>
174 Some((name, symbols, antiquotes, delimited))
182 /* external resources */
185 val Path = new Markup_String(PATH, NAME)
188 val Url = new Markup_String(URL, NAME)
191 val Doc = new Markup_String(DOC, NAME)
194 /* pretty printing */
196 val Consistent = new Properties.Boolean("consistent")
197 val Indent = new Properties.Int("indent")
198 val Width = new Properties.Int("width")
203 def apply(c: Boolean, i: Int): Markup =
205 (if (c) Consistent(c) else Nil) :::
206 (if (i != 0) Indent(i) else Nil))
207 def unapply(markup: Markup): Option[(Boolean, Int)] =
208 if (markup.name == name) {
209 val c = Consistent.unapply(markup.properties).getOrElse(false)
210 val i = Indent.unapply(markup.properties).getOrElse(0)
219 def apply(w: Int, i: Int): Markup =
221 (if (w != 0) Width(w) else Nil) :::
222 (if (i != 0) Indent(i) else Nil))
223 def unapply(markup: Markup): Option[(Int, Int)] =
224 if (markup.name == name) {
225 val w = Width.unapply(markup.properties).getOrElse(0)
226 val i = Indent.unapply(markup.properties).getOrElse(0)
233 val BULLET = "bullet"
235 val SEPARATOR = "separator"
238 /* text properties */
242 val HIDDEN = "hidden"
248 val TYPE_NAME = "type_name"
251 val CONSTANT = "constant"
252 val DYNAMIC_FACT = "dynamic_fact"
260 val SKOLEM = "skolem"
263 val NUMERAL = "numeral"
264 val LITERAL = "literal"
265 val DELIMITER = "delimiter"
266 val INNER_STRING = "inner_string"
267 val INNER_CARTOUCHE = "inner_cartouche"
268 val INNER_COMMENT = "inner_comment"
270 val TOKEN_RANGE = "token_range"
272 val SORTING = "sorting"
273 val TYPING = "typing"
274 val CLASS_PARAMETER = "class_parameter"
276 val ATTRIBUTE = "attribute"
277 val METHOD = "method"
282 val ANTIQUOTED = "antiquoted"
283 val ANTIQUOTE = "antiquote"
285 val ML_ANTIQUOTATION = "ML_antiquotation"
286 val DOCUMENT_ANTIQUOTATION = "document_antiquotation"
287 val DOCUMENT_ANTIQUOTATION_OPTION = "document_antiquotation_option"
292 val PARAGRAPH = "paragraph"
293 val TEXT_FOLD = "text_fold"
296 /* Markdown document structure */
298 val MARKDOWN_PARAGRAPH = "markdown_paragraph"
299 val MARKDOWN_ITEM = "markdown_item"
300 val Markdown_List = new Markup_String("markdown_list", "kind")
301 val Markdown_Bullet = new Markup_Int("markdown_bullet", "depth")
306 val ML_KEYWORD1 = "ML_keyword1"
307 val ML_KEYWORD2 = "ML_keyword2"
308 val ML_KEYWORD3 = "ML_keyword3"
309 val ML_DELIMITER = "ML_delimiter"
310 val ML_TVAR = "ML_tvar"
311 val ML_NUMERAL = "ML_numeral"
312 val ML_CHAR = "ML_char"
313 val ML_STRING = "ML_string"
314 val ML_COMMENT = "ML_comment"
315 val SML_STRING = "SML_string"
316 val SML_COMMENT = "SML_comment"
318 val ML_DEF = "ML_def"
319 val ML_OPEN = "ML_open"
320 val ML_STRUCTURE = "ML_structure"
321 val ML_TYPING = "ML_typing"
323 val ML_BREAKPOINT = "ML_breakpoint"
328 val COMMAND = "command"
329 val KEYWORD = "keyword"
330 val KEYWORD1 = "keyword1"
331 val KEYWORD2 = "keyword2"
332 val KEYWORD3 = "keyword3"
333 val QUASI_KEYWORD = "quasi_keyword"
334 val IMPROPER = "improper"
335 val OPERATOR = "operator"
336 val STRING = "string"
337 val ALT_STRING = "alt_string"
338 val VERBATIM = "verbatim"
339 val CARTOUCHE = "cartouche"
340 val COMMENT = "comment"
345 val Elapsed = new Properties.Double("elapsed")
346 val CPU = new Properties.Double("cpu")
347 val GC = new Properties.Double("gc")
349 object Timing_Properties
351 def apply(timing: isabelle.Timing): Properties.T =
352 Elapsed(timing.elapsed.seconds) ::: CPU(timing.cpu.seconds) ::: GC(timing.gc.seconds)
354 def unapply(props: Properties.T): Option[isabelle.Timing] =
355 (props, props, props) match {
356 case (Elapsed(elapsed), CPU(cpu), GC(gc)) =>
357 Some(new isabelle.Timing(Time.seconds(elapsed), Time.seconds(cpu), Time.seconds(gc)))
362 val TIMING = "timing"
366 def apply(timing: isabelle.Timing): Markup = Markup(TIMING, Timing_Properties(timing))
368 def unapply(markup: Markup): Option[isabelle.Timing] =
370 case Markup(TIMING, Timing_Properties(timing)) => Some(timing)
378 val Return_Code = new Properties.Int("return_code")
380 object Process_Result
382 def apply(result: Process_Result): Properties.T =
383 Return_Code(result.rc) :::
384 (if (result.timing.is_zero) Nil else Timing_Properties(result.timing))
386 def unapply(props: Properties.T): Option[Process_Result] =
388 case Return_Code(rc) =>
389 val timing = Timing_Properties.unapply(props).getOrElse(isabelle.Timing.zero)
390 Some(isabelle.Process_Result(rc, timing = timing))
396 /* command indentation */
398 object Command_Indent
400 val name = "command_indent"
401 def unapply(markup: Markup): Option[Int] =
402 if (markup.name == name) Indent.unapply(markup.properties) else None
409 val SUBGOAL = "subgoal"
416 val ACCEPTED = "accepted"
417 val FORKED = "forked"
418 val JOINED = "joined"
419 val RUNNING = "running"
420 val FINISHED = "finished"
421 val FAILED = "failed"
422 val CONSOLIDATED = "consolidated"
425 /* interactive documents */
427 val VERSION = "version"
428 val ASSIGN = "assign"
433 val PROVER_COMMAND = "prover_command"
434 val PROVER_ARG = "prover_arg"
440 val STATUS = "status"
441 val REPORT = "report"
442 val RESULT = "result"
443 val WRITELN = "writeln"
445 val INFORMATION = "information"
446 val TRACING = "tracing"
447 val WARNING = "warning"
448 val LEGACY = "legacy"
450 val PROTOCOL = "protocol"
451 val SYSTEM = "system"
452 val STDOUT = "stdout"
453 val STDERR = "stderr"
456 val WRITELN_MESSAGE = "writeln_message"
457 val STATE_MESSAGE = "state_message"
458 val INFORMATION_MESSAGE = "information_message"
459 val TRACING_MESSAGE = "tracing_message"
460 val WARNING_MESSAGE = "warning_message"
461 val LEGACY_MESSAGE = "legacy_message"
462 val ERROR_MESSAGE = "error_message"
465 WRITELN -> WRITELN_MESSAGE,
466 STATE -> STATE_MESSAGE,
467 INFORMATION -> INFORMATION_MESSAGE,
468 TRACING -> TRACING_MESSAGE,
469 WARNING -> WARNING_MESSAGE,
470 LEGACY -> LEGACY_MESSAGE,
471 ERROR -> ERROR_MESSAGE)
473 val message: String => String = messages.withDefault((s: String) => s)
475 val NO_REPORT = "no_report"
479 val INTENSIFY = "intensify"
484 val BROWSER = "browser"
485 val GRAPHVIEW = "graphview"
487 val SENDBACK = "sendback"
488 val PADDING = "padding"
489 val PADDING_LINE = (PADDING, "line")
490 val PADDING_COMMAND = (PADDING, "command")
492 val DIALOG = "dialog"
493 val Result = new Properties.String(RESULT)
495 val JEDIT_ACTION = "jedit_action"
498 /* protocol message functions */
500 val FUNCTION = "function"
501 val Function = new Properties.String(FUNCTION)
503 val COMMAND_TIMING: Properties.Entry = (FUNCTION, "command_timing")
504 val THEORY_TIMING: Properties.Entry = (FUNCTION, "theory_timing")
506 val Assign_Update: Properties.T = List((FUNCTION, "assign_update"))
507 val Removed_Versions: Properties.T = List((FUNCTION, "removed_versions"))
509 object Protocol_Handler
511 def unapply(props: Properties.T): Option[(String)] =
513 case List((FUNCTION, "protocol_handler"), (NAME, name)) => Some(name)
518 val INVOKE_SCALA = "invoke_scala"
521 def unapply(props: Properties.T): Option[(String, String)] =
523 case List((FUNCTION, INVOKE_SCALA), (NAME, name), (ID, id)) => Some((name, id))
528 val CANCEL_SCALA = "cancel_scala"
531 def unapply(props: Properties.T): Option[String] =
533 case List((FUNCTION, CANCEL_SCALA), (ID, id)) => Some(id)
540 def unapply(props: Properties.T): Option[Properties.T] =
542 case (FUNCTION, "ML_statistics") :: props => Some(props)
547 object Task_Statistics
549 def unapply(props: Properties.T): Option[Properties.T] =
551 case (FUNCTION, "task_statistics") :: props => Some(props)
556 val LOADING_THEORY = "loading_theory"
557 object Loading_Theory
559 def unapply(props: Properties.T): Option[String] =
561 case List((FUNCTION, LOADING_THEORY), (NAME, name)) => Some(name)
566 val BUILD_SESSION_FINISHED = "build_session_finished"
567 val Build_Session_Finished: Properties.T = List((FUNCTION, BUILD_SESSION_FINISHED))
569 val PRINT_OPERATIONS = "print_operations"
572 /* debugger output */
574 val DEBUGGER_STATE = "debugger_state"
575 object Debugger_State
577 def unapply(props: Properties.T): Option[String] =
579 case List((FUNCTION, DEBUGGER_STATE), (NAME, name)) => Some(name)
584 val DEBUGGER_OUTPUT = "debugger_output"
585 object Debugger_Output
587 def unapply(props: Properties.T): Option[String] =
589 case List((FUNCTION, DEBUGGER_OUTPUT), (NAME, name)) => Some(name)
595 /* simplifier trace */
597 val SIMP_TRACE_PANEL = "simp_trace_panel"
599 val SIMP_TRACE_LOG = "simp_trace_log"
600 val SIMP_TRACE_STEP = "simp_trace_step"
601 val SIMP_TRACE_RECURSE = "simp_trace_recurse"
602 val SIMP_TRACE_HINT = "simp_trace_hint"
603 val SIMP_TRACE_IGNORE = "simp_trace_ignore"
605 val SIMP_TRACE_CANCEL = "simp_trace_cancel"
606 object Simp_Trace_Cancel
608 def unapply(props: Properties.T): Option[Long] =
610 case (FUNCTION, SIMP_TRACE_CANCEL) :: Serial(i) => Some(i)
616 /* XML data representation */
618 def encode: XML.Encode.T[Markup] = (markup: Markup) =>
621 pair(string, properties)((markup.name, markup.properties))
624 def decode: XML.Decode.T[Markup] = (body: XML.Body) =>
627 val (name, props) = pair(string, properties)(body)
633 sealed case class Markup(name: String, properties: Properties.T)
635 def markup(s: String): String =
636 YXML.string_of_tree(XML.Elem(this, List(XML.Text(s))))
638 def update_properties(more_props: Properties.T): Markup =
639 if (more_props.isEmpty) this
640 else Markup(name, (more_props :\ properties) { case (p, ps) => Properties.put(ps, p) })
642 def + (entry: Properties.Entry): Markup =
643 Markup(name, Properties.put(properties, entry))