added antiquotation @{doc}, e.g. useful for demonstration purposes;
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"
94 def unapply(markup: Markup): Option[(String, String)] =
96 case Markup(ENTITY, props) =>
97 (props, props) match {
98 case (Kind(kind), Name(name)) => Some((kind, name))
108 val COMPLETION = "completion"
109 val NO_COMPLETION = "no_completion"
115 val END_LINE = "line"
116 val OFFSET = "offset"
117 val END_OFFSET = "end_offset"
121 val DEF_LINE = "def_line"
122 val DEF_OFFSET = "def_offset"
123 val DEF_END_OFFSET = "def_end_offset"
124 val DEF_FILE = "def_file"
125 val DEF_ID = "def_id"
127 val POSITION_PROPERTIES = Set(LINE, OFFSET, END_OFFSET, FILE, ID)
128 val POSITION = "position"
133 val EXPRESSION = "expression"
138 val CITATION = "citation"
139 val Citation = new Markup_String(CITATION, NAME)
142 /* embedded languages */
144 val Symbols = new Properties.Boolean("symbols")
145 val Antiquotes = new Properties.Boolean("antiquotes")
146 val Delimited = new Properties.Boolean("delimited")
148 val LANGUAGE = "language"
154 val UNKNOWN = "unknown"
156 def unapply(markup: Markup): Option[(String, Boolean, Boolean, Boolean)] =
158 case Markup(LANGUAGE, props) =>
159 (props, props, props, props) match {
160 case (Name(name), Symbols(symbols), Antiquotes(antiquotes), Delimited(delimited)) =>
161 Some((name, symbols, antiquotes, delimited))
169 /* external resources */
172 val Path = new Markup_String(PATH, NAME)
175 val Url = new Markup_String(URL, NAME)
178 val Doc = new Markup_String(DOC, NAME)
181 /* pretty printing */
183 val Block = new Markup_Int("block", "indent")
184 val Break = new Markup_Int("break", "width")
187 val BULLET = "bullet"
189 val SEPARATOR = "separator"
192 /* text properties */
196 val HIDDEN = "hidden"
202 val TYPE_NAME = "type_name"
205 val CONSTANT = "constant"
206 val DYNAMIC_FACT = "dynamic_fact"
214 val SKOLEM = "skolem"
217 val NUMERAL = "numeral"
218 val LITERAL = "literal"
219 val DELIMITER = "delimiter"
220 val INNER_STRING = "inner_string"
221 val INNER_CARTOUCHE = "inner_cartouche"
222 val INNER_COMMENT = "inner_comment"
224 val TOKEN_RANGE = "token_range"
226 val SORTING = "sorting"
227 val TYPING = "typing"
229 val ATTRIBUTE = "attribute"
230 val METHOD = "method"
235 val ANTIQUOTED = "antiquoted"
236 val ANTIQUOTE = "antiquote"
238 val ML_ANTIQUOTATION = "ML_antiquotation"
239 val DOCUMENT_ANTIQUOTATION = "document_antiquotation"
240 val DOCUMENT_ANTIQUOTATION_OPTION = "document_antiquotation_option"
245 val PARAGRAPH = "paragraph"
246 val TEXT_FOLD = "text_fold"
249 /* Markdown document structure */
251 val MARKDOWN_PARAGRAPH = "markdown_paragraph"
252 val Markdown_List = new Markup_String("markdown_list", "kind")
253 val Markdown_Item = new Markup_Int("markdown_item", "depth")
258 val ML_KEYWORD1 = "ML_keyword1"
259 val ML_KEYWORD2 = "ML_keyword2"
260 val ML_KEYWORD3 = "ML_keyword3"
261 val ML_DELIMITER = "ML_delimiter"
262 val ML_TVAR = "ML_tvar"
263 val ML_NUMERAL = "ML_numeral"
264 val ML_CHAR = "ML_char"
265 val ML_STRING = "ML_string"
266 val ML_COMMENT = "ML_comment"
267 val SML_STRING = "SML_string"
268 val SML_COMMENT = "SML_comment"
270 val ML_DEF = "ML_def"
271 val ML_OPEN = "ML_open"
272 val ML_STRUCTURE = "ML_structure"
273 val ML_TYPING = "ML_typing"
275 val ML_BREAKPOINT = "ML_breakpoint"
280 val COMMAND = "command"
281 val KEYWORD1 = "keyword1"
282 val KEYWORD2 = "keyword2"
283 val KEYWORD3 = "keyword3"
284 val QUASI_KEYWORD = "quasi_keyword"
285 val IMPROPER = "improper"
286 val OPERATOR = "operator"
287 val STRING = "string"
288 val ALT_STRING = "alt_string"
289 val VERBATIM = "verbatim"
290 val CARTOUCHE = "cartouche"
291 val COMMENT = "comment"
296 val Elapsed = new Properties.Double("elapsed")
297 val CPU = new Properties.Double("cpu")
298 val GC = new Properties.Double("gc")
300 object Timing_Properties
302 def apply(timing: isabelle.Timing): Properties.T =
303 Elapsed(timing.elapsed.seconds) ::: CPU(timing.cpu.seconds) ::: GC(timing.gc.seconds)
305 def unapply(props: Properties.T): Option[isabelle.Timing] =
306 (props, props, props) match {
307 case (Elapsed(elapsed), CPU(cpu), GC(gc)) =>
308 Some(new isabelle.Timing(Time.seconds(elapsed), Time.seconds(cpu), Time.seconds(gc)))
313 val TIMING = "timing"
317 def apply(timing: isabelle.Timing): Markup = Markup(TIMING, Timing_Properties(timing))
319 def unapply(markup: Markup): Option[isabelle.Timing] =
321 case Markup(TIMING, Timing_Properties(timing)) => Some(timing)
329 val COMMAND_TIMING = "command_timing"
334 val SUBGOALS = "subgoals"
335 val PROOF_STATE = "proof_state"
338 val SUBGOAL = "subgoal"
345 val ACCEPTED = "accepted"
346 val FORKED = "forked"
347 val JOINED = "joined"
348 val RUNNING = "running"
349 val FINISHED = "finished"
350 val FAILED = "failed"
353 /* interactive documents */
355 val VERSION = "version"
356 val ASSIGN = "assign"
361 val PROVER_COMMAND = "prover_command"
362 val PROVER_ARG = "prover_arg"
368 val STATUS = "status"
369 val REPORT = "report"
370 val RESULT = "result"
371 val WRITELN = "writeln"
373 val INFORMATION = "information"
374 val TRACING = "tracing"
375 val WARNING = "warning"
376 val LEGACY = "legacy"
378 val PROTOCOL = "protocol"
379 val SYSTEM = "system"
380 val STDOUT = "stdout"
381 val STDERR = "stderr"
384 val WRITELN_MESSAGE = "writeln_message"
385 val STATE_MESSAGE = "state_message"
386 val INFORMATION_MESSAGE = "information_message"
387 val TRACING_MESSAGE = "tracing_message"
388 val WARNING_MESSAGE = "warning_message"
389 val LEGACY_MESSAGE = "legacy_message"
390 val ERROR_MESSAGE = "error_message"
393 WRITELN -> WRITELN_MESSAGE,
394 STATE -> STATE_MESSAGE,
395 INFORMATION -> INFORMATION_MESSAGE,
396 TRACING -> TRACING_MESSAGE,
397 WARNING -> WARNING_MESSAGE,
398 LEGACY -> LEGACY_MESSAGE,
399 ERROR -> ERROR_MESSAGE)
401 val message: String => String = messages.withDefault((s: String) => s)
403 val Return_Code = new Properties.Int("return_code")
405 val NO_REPORT = "no_report"
409 val INTENSIFY = "intensify"
414 val BROWSER = "browser"
415 val GRAPHVIEW = "graphview"
417 val SENDBACK = "sendback"
418 val PADDING = "padding"
419 val PADDING_LINE = (PADDING, "line")
420 val PADDING_COMMAND = (PADDING, "command")
422 val DIALOG = "dialog"
423 val Result = new Properties.String(RESULT)
426 /* protocol message functions */
428 val FUNCTION = "function"
429 val Function = new Properties.String(FUNCTION)
431 val Assign_Update: Properties.T = List((FUNCTION, "assign_update"))
432 val Removed_Versions: Properties.T = List((FUNCTION, "removed_versions"))
434 object Protocol_Handler
436 def unapply(props: Properties.T): Option[(String)] =
438 case List((FUNCTION, "protocol_handler"), (NAME, name)) => Some(name)
443 val INVOKE_SCALA = "invoke_scala"
446 def unapply(props: Properties.T): Option[(String, String)] =
448 case List((FUNCTION, INVOKE_SCALA), (NAME, name), (ID, id)) => Some((name, id))
453 val CANCEL_SCALA = "cancel_scala"
456 def unapply(props: Properties.T): Option[String] =
458 case List((FUNCTION, CANCEL_SCALA), (ID, id)) => Some(id)
465 def unapply(props: Properties.T): Option[Properties.T] =
467 case (FUNCTION, "ML_statistics") :: stats => Some(stats)
472 object Task_Statistics
474 def unapply(props: Properties.T): Option[Properties.T] =
476 case (FUNCTION, "task_statistics") :: stats => Some(stats)
481 val LOADING_THEORY = "loading_theory"
482 object Loading_Theory
484 def unapply(props: Properties.T): Option[String] =
486 case List((FUNCTION, LOADING_THEORY), (NAME, name)) => Some(name)
491 val BUILD_THEORIES_RESULT = "build_theories_result"
492 object Build_Theories_Result
494 def unapply(props: Properties.T): Option[String] =
496 case List((FUNCTION, BUILD_THEORIES_RESULT), (ID, id)) => Some(id)
501 val PRINT_OPERATIONS = "print_operations"
504 /* debugger output */
506 val DEBUGGER_STATE = "debugger_state"
507 object Debugger_State
509 def unapply(props: Properties.T): Option[String] =
511 case List((FUNCTION, DEBUGGER_STATE), (NAME, name)) => Some(name)
516 val DEBUGGER_OUTPUT = "debugger_output"
517 object Debugger_Output
519 def unapply(props: Properties.T): Option[String] =
521 case List((FUNCTION, DEBUGGER_OUTPUT), (NAME, name)) => Some(name)
527 /* simplifier trace */
529 val SIMP_TRACE_PANEL = "simp_trace_panel"
531 val SIMP_TRACE_LOG = "simp_trace_log"
532 val SIMP_TRACE_STEP = "simp_trace_step"
533 val SIMP_TRACE_RECURSE = "simp_trace_recurse"
534 val SIMP_TRACE_HINT = "simp_trace_hint"
535 val SIMP_TRACE_IGNORE = "simp_trace_ignore"
537 val SIMP_TRACE_CANCEL = "simp_trace_cancel"
538 object Simp_Trace_Cancel
540 def unapply(props: Properties.T): Option[Long] =
542 case (FUNCTION, SIMP_TRACE_CANCEL) :: Serial(i) => Some(i)
549 sealed case class Markup(name: String, properties: Properties.T)
551 def markup(s: String): String =
552 YXML.string_of_tree(XML.Elem(this, List(XML.Text(s))))