diff --git a/src/Pure/General/html.scala b/src/Pure/General/html.scala --- a/src/Pure/General/html.scala +++ b/src/Pure/General/html.scala @@ -1,462 +1,460 @@ /* Title: Pure/General/html.scala Author: Makarius HTML presentation elements. */ package isabelle import org.jsoup.nodes.{Entities => Jsoup_Entities, Document => Jsoup_Document} import org.jsoup.Jsoup object HTML { /* attributes */ class Attribute(val name: String, value: String) { def xml: XML.Attribute = name -> value def apply(elem: XML.Elem): XML.Elem = elem + xml } def id(s: String): Attribute = new Attribute("id", s) def class_(name: String): Attribute = new Attribute("class", name) def width(w: Int): Attribute = new Attribute("width", w.toString) def height(h: Int): Attribute = new Attribute("height", h.toString) def size(w: Int, h: Int)(elem: XML.Elem): XML.Elem = width(w)(height(h)(elem)) val entity_def: Attribute = class_("entity_def") val entity_ref: Attribute = class_("entity_ref") /* structured markup operators */ def text(txt: String): XML.Body = if (txt.isEmpty) Nil else List(XML.Text(txt)) val break: XML.Body = List(XML.elem("br")) val nl: XML.Body = List(XML.Text("\n")) class Operator(val name: String) { def apply(body: XML.Body): XML.Elem = XML.elem(name, body) def apply(att: Attribute, body: XML.Body): XML.Elem = att(apply(body)) def apply(c: String, body: XML.Body): XML.Elem = apply(class_(c), body) } class Heading(name: String) extends Operator(name) { def apply(txt: String): XML.Elem = super.apply(text(txt)) def apply(att: Attribute, txt: String): XML.Elem = super.apply(att, text(txt)) def apply(c: String, txt: String): XML.Elem = super.apply(c, text(txt)) } val div = new Operator("div") val span = new Operator("span") val pre = new Operator("pre") val par = new Operator("p") val sub = new Operator("sub") val sup = new Operator("sup") val emph = new Operator("em") val bold = new Operator("b") val code = new Operator("code") val item = new Operator("li") val list = new Operator("ul") val `enum` = new Operator("ol") val descr = new Operator("dl") val dt = new Operator("dt") val dd = new Operator("dd") val title = new Heading("title") val chapter = new Heading("h1") val section = new Heading("h2") val subsection = new Heading("h3") val subsubsection = new Heading("h4") val paragraph = new Heading("h5") val subparagraph = new Heading("h6") def itemize(items: List[XML.Body]): XML.Elem = list(items.map(item(_))) def enumerate(items: List[XML.Body]): XML.Elem = `enum`(items.map(item(_))) def description(items: List[(XML.Body, XML.Body)]): XML.Elem = descr(items.flatMap({ case (x, y) => List(dt(x), dd(y)) })) def link(href: String, body: XML.Body): XML.Elem = XML.Elem(Markup("a", List("href" -> href)), if (body.isEmpty) text(href) else body) def link(path: Path, body: XML.Body): XML.Elem = link(path.implode, body) def image(src: String, alt: String = ""): XML.Elem = XML.Elem(Markup("img", List("src" -> src) ::: proper_string(alt).map("alt" -> _).toList), Nil) def source(body: XML.Body): XML.Elem = pre("source", body) def source(src: String): XML.Elem = source(text(src)) def style(s: String): XML.Elem = XML.elem("style", text(s)) def style_file(href: String): XML.Elem = XML.Elem(Markup("link", List("rel" -> "stylesheet", "type" -> "text/css", "href" -> href)), Nil) def style_file(path: Path): XML.Elem = style_file(Url.print_file(path.file)) def script(s: String): XML.Elem = XML.elem("script", text(s)) def script_file(href: String): XML.Elem = XML.Elem(Markup("script", List("src" -> href)), Nil) def script_file(path: Path): XML.Elem = script_file(Url.print_file(path.file)) /* href */ def relative_href(location: Path, base: Option[Path] = None): String = { val path = base match { case None => val path = location.expand if (path.is_absolute) Exn.error("Relative href location expected: " + path) else path case Some(base_dir) => val path1 = base_dir.absolute_file.toPath val path2 = location.absolute_file.toPath try { File.path(path1.relativize(path2).toFile) } catch { case _: IllegalArgumentException => Exn.error("Failed to relativize href location " + path2 + " with wrt. base " + path1) } } if (path.is_current) "" else path.implode } /* output text with control symbols */ private val control: Map[Symbol.Symbol, Operator] = Map( Symbol.sub -> sub, Symbol.sub_decoded -> sub, Symbol.sup -> sup, Symbol.sup_decoded -> sup, Symbol.bold -> bold, Symbol.bold_decoded -> bold) private val control_block_begin: Map[Symbol.Symbol, Operator] = Map( Symbol.bsub -> sub, Symbol.bsub_decoded -> sub, Symbol.bsup -> sup, Symbol.bsup_decoded -> sup) private val control_block_end: Map[Symbol.Symbol, Operator] = Map( Symbol.esub -> sub, Symbol.esub_decoded -> sub, Symbol.esup -> sup, Symbol.esup_decoded -> sup) def is_control(sym: Symbol.Symbol): Boolean = control.isDefinedAt(sym) def is_control_block_begin(sym: Symbol.Symbol): Boolean = control_block_begin.isDefinedAt(sym) def is_control_block_end(sym: Symbol.Symbol): Boolean = control_block_end.isDefinedAt(sym) def is_control_block_pair(bg: Symbol.Symbol, en: Symbol.Symbol): Boolean = { val bg_decoded = Symbol.decode(bg) val en_decoded = Symbol.decode(en) bg_decoded == Symbol.bsub_decoded && en_decoded == Symbol.esub_decoded || bg_decoded == Symbol.bsup_decoded && en_decoded == Symbol.esup_decoded } def check_control_blocks(body: XML.Body): Boolean = { var ok = true var open = List.empty[Symbol.Symbol] for { case XML.Text(text) <- body; sym <- Symbol.iterator(text) } { if (is_control_block_begin(sym)) open ::= sym else if (is_control_block_end(sym)) { open match { case bg :: rest if is_control_block_pair(bg, sym) => open = rest case _ => ok = false } } } ok && open.isEmpty } def output( s: StringBuilder, text: String, control_blocks: Boolean, hidden: Boolean, permissive: Boolean ): Unit = { def output_string(str: String): Unit = XML.output_string(s, str, permissive = permissive) def output_hidden(body: => Unit): Unit = if (hidden) { s ++= ""; body; s ++= "" } def output_symbol(sym: Symbol.Symbol): Unit = if (sym != "") { control_block_begin.get(sym) match { case Some(op) if control_blocks => output_hidden(output_string(sym)) XML.output_elem(s, Markup(op.name, Nil)) case _ => control_block_end.get(sym) match { case Some(op) if control_blocks => XML.output_elem_end(s, op.name) output_hidden(output_string(sym)) case _ => if (hidden && Symbol.is_control_encoded(sym)) { output_hidden(output_string(Symbol.control_prefix)) s ++= "" output_string(Symbol.control_name(sym).get) s ++= "" output_hidden(output_string(Symbol.control_suffix)) } else output_string(sym) } } } var ctrl = "" for (sym <- Symbol.iterator(text)) { if (is_control(sym)) { output_symbol(ctrl); ctrl = sym } else { control.get(ctrl) match { case Some(op) if Symbol.is_controllable(sym) => output_hidden(output_symbol(ctrl)) XML.output_elem(s, Markup(op.name, Nil)) output_symbol(sym) XML.output_elem_end(s, op.name) case _ => output_symbol(ctrl) output_symbol(sym) } ctrl = "" } } output_symbol(ctrl) } def output(text: String): String = { val control_blocks = check_control_blocks(List(XML.Text(text))) Library.make_string(output(_, text, control_blocks = control_blocks, hidden = false, permissive = true)) } /* output XML as HTML */ private val structural_elements = Set("head", "body", "meta", "div", "pre", "p", "title", "h1", "h2", "h3", "h4", "h5", "h6", "ul", "ol", "dl", "li", "dt", "dd") def output(s: StringBuilder, xml: XML.Body, hidden: Boolean, structural: Boolean): Unit = { def output_body(body: XML.Body): Unit = { val control_blocks = check_control_blocks(body) body foreach { case XML.Elem(markup, Nil) => XML.output_elem(s, markup, end = true) - case XML.Elem(Markup(RAW, _), List(XML.Text(raw))) => + case XML.Elem(Markup(Markup.RAW_HTML, _), List(XML.Text(raw))) => s ++= raw case XML.Elem(markup, ts) => if (structural && structural_elements(markup.name)) s += '\n' XML.output_elem(s, markup) output_body(ts) XML.output_elem_end(s, markup.name) if (structural && structural_elements(markup.name)) s += '\n' case XML.Text(txt) => output(s, txt, control_blocks = control_blocks, hidden = hidden, permissive = true) } } output_body(xml) } def output(body: XML.Body, hidden: Boolean, structural: Boolean): String = Library.make_string(output(_, body, hidden, structural), capacity = XML.text_length(body) * 2) def output(tree: XML.Tree, hidden: Boolean, structural: Boolean): String = output(List(tree), hidden, structural) /* input */ - val RAW = "raw" - def input(text: String): String = Jsoup_Entities.unescape(text) - def input_raw(text: String): XML.Elem = XML.elem(RAW, List(XML.Text(input(text)))) + def input_raw(text: String): XML.Elem = XML.elem(Markup.RAW_HTML, List(XML.Text(input(text)))) def parse_document(html: String): Jsoup_Document = Jsoup.parse(html) def get_document(url: String): Jsoup_Document = Jsoup.connect(url).get() /* messages */ // background val writeln_message: Attribute = class_("writeln_message") val warning_message: Attribute = class_("warning_message") val error_message: Attribute = class_("error_message") // underline val writeln: Attribute = class_("writeln") val warning: Attribute = class_("warning") val error: Attribute = class_("error") /* tooltips */ def tooltip(item: XML.Body, tip: XML.Body): XML.Elem = span(item ::: List(div("tooltip", tip))) def tooltip_errors(item: XML.Body, msgs: List[XML.Body]): XML.Elem = HTML.error(tooltip(item, msgs.map(msg => error_message(pre(msg))))) /* GUI elements */ object GUI { private def optional_value(text: String): XML.Attributes = proper_string(text).map(a => "value" -> a).toList private def optional_name(name: String): XML.Attributes = proper_string(name).map(a => "name" -> a).toList private def optional_title(tooltip: String): XML.Attributes = proper_string(tooltip).map(a => "title" -> a).toList private def optional_submit(submit: Boolean): XML.Attributes = if (submit) List("onChange" -> "this.form.submit()") else Nil private def optional_checked(selected: Boolean): XML.Attributes = if (selected) List("checked" -> "") else Nil private def optional_action(action: String): XML.Attributes = proper_string(action).map(a => "action" -> a).toList private def optional_onclick(script: String): XML.Attributes = proper_string(script).map(a => "onclick" -> a).toList private def optional_onchange(script: String): XML.Attributes = proper_string(script).map(a => "onchange" -> a).toList def button(body: XML.Body, name: String = "", tooltip: String = "", submit: Boolean = false, script: String = ""): XML.Elem = XML.Elem( Markup("button", List("type" -> (if (submit) "submit" else "button"), "value" -> "true") ::: optional_name(name) ::: optional_title(tooltip) ::: optional_onclick(script)), body) def checkbox(body: XML.Body, name: String = "", tooltip: String = "", submit: Boolean = false, selected: Boolean = false, script: String = ""): XML.Elem = XML.elem("label", XML.elem( Markup("input", List("type" -> "checkbox", "value" -> "true") ::: optional_name(name) ::: optional_title(tooltip) ::: optional_submit(submit) ::: optional_checked(selected) ::: optional_onchange(script))) :: body) def text_field(columns: Int = 0, text: String = "", name: String = "", tooltip: String = "", submit: Boolean = false, script: String = ""): XML.Elem = XML.elem(Markup("input", List("type" -> "text") ::: (if (columns > 0) List("size" -> columns.toString) else Nil) ::: optional_value(text) ::: optional_name(name) ::: optional_title(tooltip) ::: optional_submit(submit) ::: optional_onchange(script))) def parameter(text: String = "", name: String = ""): XML.Elem = XML.elem( Markup("input", List("type" -> "hidden") ::: optional_value(text) ::: optional_name(name))) def form(body: XML.Body, name: String = "", action: String = "", http_post: Boolean = false) : XML.Elem = XML.Elem( Markup("form", optional_name(name) ::: optional_action(action) ::: (if (http_post) List("method" -> "post") else Nil)), body) } /* GUI layout */ object Wrap_Panel { enum Alignment { case left, right, center } def apply( contents: List[XML.Elem], name: String = "", action: String = "", alignment: Alignment = Alignment.right ): XML.Elem = { val body = Library.separate(XML.Text(" "), contents) GUI.form(List(div(body) + ("style" -> ("text-align: " + alignment))), name = name, action = action) } } /* document */ val header: String = XML.header + """ """ val footer: String = """""" val head_meta: XML.Elem = XML.Elem(Markup("meta", List("http-equiv" -> "Content-Type", "content" -> "text/html; charset=utf-8")), Nil) def output_document( head: XML.Body, body: XML.Body, css: String = "isabelle.css", hidden: Boolean = true, structural: Boolean = true ): String = { cat_lines( List( header, output( XML.elem("head", head_meta :: (if (css == "") Nil else List(style_file(css))) ::: head), hidden = hidden, structural = structural), output(XML.elem("body", body), hidden = hidden, structural = structural), footer)) } /* fonts */ val fonts_path: Path = Path.explode("fonts") def init_fonts(dir: Path): Unit = { val fonts_dir = Isabelle_System.make_directory(dir + HTML.fonts_path) for (entry <- Isabelle_Fonts.fonts(hidden = true)) Isabelle_System.copy_file(entry.path, fonts_dir) } def fonts_dir(prefix: String)(ttf_name: String): String = prefix + "/" + ttf_name def fonts_url(): String => String = (for (entry <- Isabelle_Fonts.fonts(hidden = true)) yield (entry.path.file_name -> Url.print_file(entry.path.file))).toMap def fonts_css(make_url: String => String = fonts_url()): String = { def font_face(entry: Isabelle_Fonts.Entry): String = cat_lines( List( "@font-face {", " font-family: '" + entry.family + "';", " src: url('" + make_url(entry.path.file_name) + "') format('truetype');") ::: (if (entry.is_bold) List(" font-weight: bold;") else Nil) ::: (if (entry.is_italic) List(" font-style: italic;") else Nil) ::: List("}")) ("/* Isabelle fonts */" :: Isabelle_Fonts.fonts(hidden = true).map(font_face)) .mkString("", "\n\n", "\n") } def fonts_css_dir(prefix: String = ""): String = fonts_css(fonts_dir(Url.append_path(prefix, fonts_path.implode))) /* document directory context (fonts + css) */ def isabelle_css: Path = Path.explode("~~/etc/isabelle.css") def write_document(base_dir: Path, name: String, head: XML.Body, body: XML.Body, root: Option[Path] = None, css: String = isabelle_css.file_name, hidden: Boolean = true, structural: Boolean = true ): Unit = { Isabelle_System.make_directory(base_dir) val fonts_prefix = relative_href(root getOrElse base_dir, base = Some(base_dir)) val fonts = fonts_css_dir(fonts_prefix) File.write(base_dir + isabelle_css.base, fonts + "\n\n" + File.read(isabelle_css)) File.write(base_dir + Path.basic(name), output_document(head, body, css = css, hidden = hidden, structural = structural)) } } diff --git a/src/Pure/PIDE/markup.scala b/src/Pure/PIDE/markup.scala --- a/src/Pure/PIDE/markup.scala +++ b/src/Pure/PIDE/markup.scala @@ -1,790 +1,795 @@ /* Title: Pure/PIDE/markup.scala Author: Makarius Quasi-abstract markup elements. */ package isabelle object Markup { /* elements */ object Elements { def apply(elems: Set[String]): Elements = new Elements(elems) def apply(elems: String*): Elements = apply(Set(elems: _*)) val empty: Elements = apply() val full: Elements = new Elements(Set.empty) { override def apply(elem: String): Boolean = true override def toString: String = "Elements.full" } } sealed class Elements private[Markup](private val rep: Set[String]) { def apply(elem: String): Boolean = rep.contains(elem) def + (elem: String): Elements = new Elements(rep + elem) def ++ (elems: Elements): Elements = new Elements(rep ++ elems.rep) def - (elem: String): Elements = new Elements(rep - elem) def -- (elems: Elements): Elements = new Elements(rep -- elems.rep) override def toString: String = rep.mkString("Elements(", ",", ")") } /* properties */ val NAME = "name" val Name = new Properties.String(NAME) val XNAME = "xname" val XName = new Properties.String(XNAME) val KIND = "kind" val Kind = new Properties.String(KIND) val CONTENT = "content" val Content = new Properties.String(CONTENT) val SERIAL = "serial" val Serial = new Properties.Long(SERIAL) val INSTANCE = "instance" val Instance = new Properties.String(INSTANCE) /* basic markup */ val Empty: Markup = Markup("", Nil) val Broken: Markup = Markup("broken", Nil) class Markup_Elem(val name: String) { def apply(props: Properties.T = Nil): Markup = Markup(name, props) def unapply(markup: Markup): Option[Properties.T] = if (markup.name == name) Some(markup.properties) else None } class Markup_String(val name: String, prop: String) { val Prop: Properties.String = new Properties.String(prop) def apply(s: String): Markup = Markup(name, Prop(s)) def unapply(markup: Markup): Option[String] = if (markup.name == name) Prop.unapply(markup.properties) else None def get(markup: Markup): String = unapply(markup).getOrElse("") } class Markup_Int(val name: String, prop: String) { val Prop: Properties.Int = new Properties.Int(prop) def apply(i: Int): Markup = Markup(name, Prop(i)) def unapply(markup: Markup): Option[Int] = if (markup.name == name) Prop.unapply(markup.properties) else None def get(markup: Markup): Int = unapply(markup).getOrElse(0) } class Markup_Long(val name: String, prop: String) { val Prop: Properties.Long = new Properties.Long(prop) def apply(i: Long): Markup = Markup(name, Prop(i)) def unapply(markup: Markup): Option[Long] = if (markup.name == name) Prop.unapply(markup.properties) else None def get(markup: Markup): Long = unapply(markup).getOrElse(0) } /* meta data */ val META_TITLE = "meta_title" val META_CREATOR = "meta_creator" val META_CONTRIBUTOR = "meta_contributor" val META_DATE = "meta_date" val META_LICENSE = "meta_license" val META_DESCRIPTION = "meta_description" /* formal entities */ val BINDING = "binding" val ENTITY = "entity" object Entity { val Def = new Markup_Long(ENTITY, "def") val Ref = new Markup_Long(ENTITY, "ref") object Occ { def unapply(markup: Markup): Option[Long] = Def.unapply(markup) orElse Ref.unapply(markup) } def unapply(markup: Markup): Option[(String, String)] = markup match { case Markup(ENTITY, props) => Some((Kind.get(props), Name.get(props))) case _ => None } } /* completion */ val COMPLETION = "completion" val NO_COMPLETION = "no_completion" val UPDATE = "update" /* position */ val LINE = "line" val END_LINE = "line" val OFFSET = "offset" val END_OFFSET = "end_offset" val LABEL = "label" val FILE = "file" val ID = "id" val DEF_LINE = "def_line" val DEF_OFFSET = "def_offset" val DEF_END_OFFSET = "def_end_offset" val DEF_LABEL = "def_label" val DEF_FILE = "def_file" val DEF_ID = "def_id" val POSITION = "position" val POSITION_PROPERTIES = Set(LINE, OFFSET, END_OFFSET, LABEL, FILE, ID) def position_property(entry: Properties.Entry): Boolean = POSITION_PROPERTIES(entry._1) /* position "def" name */ private val def_names: Map[String, String] = Map(LINE -> DEF_LINE, OFFSET -> DEF_OFFSET, END_OFFSET -> DEF_END_OFFSET, LABEL -> DEF_LABEL, FILE -> DEF_FILE, ID -> DEF_ID) def def_name(a: String): String = def_names.getOrElse(a, a + "def_") /* expression */ val EXPRESSION = "expression" object Expression { def unapply(markup: Markup): Option[String] = markup match { case Markup(EXPRESSION, props) => Some(Kind.get(props)) case _ => None } } /* citation */ val CITATION = "citation" val Citation = new Markup_String(CITATION, NAME) /* embedded languages */ val Symbols = new Properties.Boolean("symbols") val Antiquotes = new Properties.Boolean("antiquotes") val Delimited = new Properties.Boolean("delimited") val LANGUAGE = "language" object Language { val DOCUMENT = "document" val ANTIQUOTATION = "antiquotation" val ML = "ML" val SML = "SML" val PATH = "path" val UNKNOWN = "unknown" def apply(name: String, symbols: Boolean, antiquotes: Boolean, delimited: Boolean): Language = new Language(name, symbols, antiquotes, delimited) val outer: Language = apply("", true, false, false) def unapply(markup: Markup): Option[Language] = markup match { case Markup(LANGUAGE, props) => (props, props, props, props) match { case (Name(name), Symbols(symbols), Antiquotes(antiquotes), Delimited(delimited)) => Some(apply(name, symbols, antiquotes, delimited)) case _ => None } case _ => None } } class Language private( val name: String, val symbols: Boolean, val antiquotes: Boolean, val delimited: Boolean ) { override def toString: String = name def is_document: Boolean = name == Language.DOCUMENT def is_antiquotation: Boolean = name == Language.ANTIQUOTATION def is_path: Boolean = name == Language.PATH def description: String = Word.implode(Word.explode('_', name)) } /* external resources */ val PATH = "path" val Path = new Markup_String(PATH, NAME) val EXPORT_PATH = "export_path" val Export_Path = new Markup_String(EXPORT_PATH, NAME) val URL = "url" val Url = new Markup_String(URL, NAME) val DOC = "doc" val Doc = new Markup_String(DOC, NAME) /* pretty printing */ val Consistent = new Properties.Boolean("consistent") val Indent = new Properties.Int("indent") val Width = new Properties.Int("width") object Block { val name = "block" def apply(c: Boolean, i: Int): Markup = Markup(name, (if (c) Consistent(c) else Nil) ::: (if (i != 0) Indent(i) else Nil)) def unapply(markup: Markup): Option[(Boolean, Int)] = if (markup.name == name) { val c = Consistent.get(markup.properties) val i = Indent.get(markup.properties) Some((c, i)) } else None } object Break { val name = "break" def apply(w: Int, i: Int): Markup = Markup(name, (if (w != 0) Width(w) else Nil) ::: (if (i != 0) Indent(i) else Nil)) def unapply(markup: Markup): Option[(Int, Int)] = if (markup.name == name) { val w = Width.get(markup.properties) val i = Indent.get(markup.properties) Some((w, i)) } else None } val ITEM = "item" val BULLET = "bullet" val SEPARATOR = "separator" /* text properties */ val WORDS = "words" val HIDDEN = "hidden" val DELETE = "delete" /* misc entities */ val SESSION = "session" val THEORY = "theory" val CLASS = "class" val LOCALE = "locale" val BUNDLE = "bundle" val TYPE_NAME = "type_name" val CONSTANT = "constant" val AXIOM = "axiom" val FACT = "fact" val ORACLE = "oracle" val FIXED = "fixed" val CASE = "case" val DYNAMIC_FACT = "dynamic_fact" val LITERAL_FACT = "literal_fact" val ATTRIBUTE = "attribute" val METHOD = "method" val METHOD_MODIFIER = "method_modifier" /* inner syntax */ val TFREE = "tfree" val TVAR = "tvar" val FREE = "free" val SKOLEM = "skolem" val BOUND = "bound" val VAR = "var" val NUMERAL = "numeral" val LITERAL = "literal" val DELIMITER = "delimiter" val INNER_STRING = "inner_string" val INNER_CARTOUCHE = "inner_cartouche" val TOKEN_RANGE = "token_range" val SORTING = "sorting" val TYPING = "typing" val CLASS_PARAMETER = "class_parameter" /* antiquotations */ val ANTIQUOTED = "antiquoted" val ANTIQUOTE = "antiquote" val ML_ANTIQUOTATION = "ML_antiquotation" val DOCUMENT_ANTIQUOTATION = "document_antiquotation" val DOCUMENT_ANTIQUOTATION_OPTION = "document_antiquotation_option" /* document text */ val RAW_TEXT = "raw_text" val PLAIN_TEXT = "plain_text" val PARAGRAPH = "paragraph" val TEXT_FOLD = "text_fold" object Document_Tag extends Markup_String("document_tag", NAME) { val IMPORTANT = "important" val UNIMPORTANT = "unimportant" } + /* HTML */ + + val RAW_HTML = "raw_html" + + /* LaTeX */ val Document_Latex = new Markup_Elem("document_latex") val Latex_Output = new Markup_Elem("latex_output") val Latex_Macro0 = new Markup_String("latex_macro0", NAME) val Latex_Macro = new Markup_String("latex_macro", NAME) val Latex_Environment = new Markup_String("latex_environment", NAME) val Latex_Heading = new Markup_String("latex_heading", KIND) val Latex_Body = new Markup_String("latex_body", KIND) val Latex_Delim = new Markup_String("latex_delim", NAME) val Latex_Tag = new Markup_String("latex_tag", NAME) val Latex_Cite = new Markup_Elem("latex_cite") val Citations = new Properties.String("citations") val Latex_Index_Item = new Markup_Elem("latex_index_item") val Latex_Index_Entry = new Markup_String("latex_index_entry", KIND) val Optional_Argument = new Properties.String("optional_argument") /* Markdown document structure */ val MARKDOWN_PARAGRAPH = "markdown_paragraph" val MARKDOWN_ITEM = "markdown_item" val Markdown_Bullet = new Markup_Int("markdown_bullet", "depth") val Markdown_List = new Markup_String("markdown_list", "kind") val ITEMIZE = "itemize" val ENUMERATE = "enumerate" val DESCRIPTION = "description" /* ML */ val ML_KEYWORD1 = "ML_keyword1" val ML_KEYWORD2 = "ML_keyword2" val ML_KEYWORD3 = "ML_keyword3" val ML_DELIMITER = "ML_delimiter" val ML_TVAR = "ML_tvar" val ML_NUMERAL = "ML_numeral" val ML_CHAR = "ML_char" val ML_STRING = "ML_string" val ML_COMMENT = "ML_comment" val ML_DEF = "ML_def" val ML_OPEN = "ML_open" val ML_STRUCTURE = "ML_structure" val ML_TYPING = "ML_typing" val ML_BREAKPOINT = "ML_breakpoint" /* outer syntax */ val COMMAND_SPAN = "command_span" object Command_Span { sealed case class Arg(name: String, kind: String) { def properties: Properties.T = (if (name.isEmpty) Nil else Name(name)) ::: (if (kind.isEmpty) Nil else Kind(kind)) } def apply(arg: Arg): Markup = Markup(COMMAND_SPAN, arg.properties) def apply(name: String, kind: String): Markup = apply(Arg(name, kind)) def unapply(markup: Markup): Option[Arg] = if (markup.name == COMMAND_SPAN) { Some(Arg(Name.get(markup.properties), Kind.get(markup.properties))) } else None } val COMMAND = "command" val KEYWORD = "keyword" val KEYWORD1 = "keyword1" val KEYWORD2 = "keyword2" val KEYWORD3 = "keyword3" val QUASI_KEYWORD = "quasi_keyword" val IMPROPER = "improper" val OPERATOR = "operator" val STRING = "string" val ALT_STRING = "alt_string" val CARTOUCHE = "cartouche" val COMMENT = "comment" val LOAD_COMMAND = "load_command" /* comments */ val COMMENT1 = "comment1" val COMMENT2 = "comment2" val COMMENT3 = "comment3" /* timing */ val Elapsed = new Properties.Double("elapsed") val CPU = new Properties.Double("cpu") val GC = new Properties.Double("gc") object Timing_Properties { def apply(timing: isabelle.Timing): Properties.T = Elapsed(timing.elapsed.seconds) ::: CPU(timing.cpu.seconds) ::: GC(timing.gc.seconds) def unapply(props: Properties.T): Option[isabelle.Timing] = (props, props, props) match { case (Elapsed(elapsed), CPU(cpu), GC(gc)) => Some(new isabelle.Timing(Time.seconds(elapsed), Time.seconds(cpu), Time.seconds(gc))) case _ => None } def get(props: Properties.T): isabelle.Timing = unapply(props).getOrElse(isabelle.Timing.zero) } val TIMING = "timing" object Timing { def apply(timing: isabelle.Timing): Markup = Markup(TIMING, Timing_Properties(timing)) def unapply(markup: Markup): Option[isabelle.Timing] = markup match { case Markup(TIMING, Timing_Properties(timing)) => Some(timing) case _ => None } } /* process result */ val Return_Code = new Properties.Int("return_code") object Process_Result { def apply(result: Process_Result): Properties.T = Return_Code(result.rc) ::: (if (result.timing.is_zero) Nil else Timing_Properties(result.timing)) def unapply(props: Properties.T): Option[Process_Result] = props match { case Return_Code(rc) => val timing = Timing_Properties.unapply(props).getOrElse(isabelle.Timing.zero) Some(isabelle.Process_Result(rc, timing = timing)) case _ => None } } /* command indentation */ val Command_Indent = new Markup_Int("command_indent", Indent.name) /* goals */ val GOAL = "goal" val SUBGOAL = "subgoal" /* command status */ val TASK = "task" val ACCEPTED = "accepted" val FORKED = "forked" val JOINED = "joined" val RUNNING = "running" val FINISHED = "finished" val FAILED = "failed" val CANCELED = "canceled" val INITIALIZED = "initialized" val FINALIZED = "finalized" val CONSOLIDATING = "consolidating" val CONSOLIDATED = "consolidated" /* interactive documents */ val VERSION = "version" val ASSIGN = "assign" /* prover process */ val PROVER_COMMAND = "prover_command" val PROVER_ARG = "prover_arg" /* messages */ val INIT = "init" val STATUS = "status" val REPORT = "report" val RESULT = "result" val WRITELN = "writeln" val STATE = "state" val INFORMATION = "information" val TRACING = "tracing" val WARNING = "warning" val LEGACY = "legacy" val ERROR = "error" val NODES_STATUS = "nodes_status" val PROTOCOL = "protocol" val SYSTEM = "system" val STDOUT = "stdout" val STDERR = "stderr" val EXIT = "exit" val WRITELN_MESSAGE = "writeln_message" val STATE_MESSAGE = "state_message" val INFORMATION_MESSAGE = "information_message" val TRACING_MESSAGE = "tracing_message" val WARNING_MESSAGE = "warning_message" val LEGACY_MESSAGE = "legacy_message" val ERROR_MESSAGE = "error_message" val messages = Map( WRITELN -> WRITELN_MESSAGE, STATE -> STATE_MESSAGE, INFORMATION -> INFORMATION_MESSAGE, TRACING -> TRACING_MESSAGE, WARNING -> WARNING_MESSAGE, LEGACY -> LEGACY_MESSAGE, ERROR -> ERROR_MESSAGE) val message: String => String = messages.withDefault((s: String) => s) val NO_REPORT = "no_report" val BAD = "bad" val INTENSIFY = "intensify" /* ML profiling */ val COUNT = "count" val ML_PROFILING_ENTRY = "ML_profiling_entry" val ML_PROFILING = "ML_profiling" object ML_Profiling { def unapply_entry(tree: XML.Tree): Option[isabelle.ML_Profiling.Entry] = tree match { case XML.Elem(Markup(ML_PROFILING_ENTRY, List((NAME, name), (COUNT, Value.Long(count)))), _) => Some(isabelle.ML_Profiling.Entry(name, count)) case _ => None } def unapply_report(tree: XML.Tree): Option[isabelle.ML_Profiling.Report] = tree match { case XML.Elem(Markup(ML_PROFILING, List((KIND, kind))), body) => Some(isabelle.ML_Profiling.Report(kind, body.flatMap(unapply_entry))) case _ => None } } /* active areas */ val BROWSER = "browser" val GRAPHVIEW = "graphview" val THEORY_EXPORTS = "theory_exports" val SENDBACK = "sendback" val PADDING = "padding" val PADDING_LINE = (PADDING, "line") val PADDING_COMMAND = (PADDING, "command") val DIALOG = "dialog" val Result = new Properties.String(RESULT) val JEDIT_ACTION = "jedit_action" /* protocol message functions */ val FUNCTION = "function" class Function(val name: String) { val THIS: Properties.Entry = (FUNCTION, name) } class Properties_Function(name: String) extends Function(name) { def unapply(props: Properties.T): Option[Properties.T] = props match { case THIS :: args => Some(args) case _ => None } } class Name_Function(name: String) extends Function(name) { def unapply(props: Properties.T): Option[String] = props match { case List(THIS, (NAME, a)) => Some(a) case _ => None } } object ML_Statistics extends Function("ML_statistics") { def unapply(props: Properties.T): Option[(Long, String)] = props match { case List(THIS, ("pid", Value.Long(pid)), ("stats_dir", stats_dir)) => Some((pid, stats_dir)) case _ => None } } val command_timing_properties: Set[String] = Set(FILE, OFFSET, NAME, Elapsed.name) def command_timing_property(entry: Properties.Entry): Boolean = command_timing_properties(entry._1) object Command_Timing extends Properties_Function("command_timing") object Theory_Timing extends Properties_Function("theory_timing") object Session_Timing extends Properties_Function("session_timing") { val Threads = new Properties.Int("threads") } object Task_Statistics extends Properties_Function("task_statistics") object Loading_Theory extends Properties_Function("loading_theory") object Build_Session_Finished extends Function("build_session_finished") object Commands_Accepted extends Function("commands_accepted") object Assign_Update extends Function("assign_update") object Removed_Versions extends Function("removed_versions") object Invoke_Scala extends Function("invoke_scala") { def unapply(props: Properties.T): Option[(String, String)] = props match { case List(THIS, (NAME, name), (ID, id)) => Some((name, id)) case _ => None } } object Cancel_Scala extends Function("cancel_scala") { def unapply(props: Properties.T): Option[String] = props match { case List(THIS, (ID, id)) => Some(id) case _ => None } } val PRINT_OPERATIONS = "print_operations" /* export */ val EXPORT = "export" val THEORY_NAME = "theory_name" val EXECUTABLE = "executable" val COMPRESS = "compress" val STRICT = "strict" /* debugger output */ val DEBUGGER_STATE = "debugger_state" object Debugger_State { def unapply(props: Properties.T): Option[String] = props match { case List((FUNCTION, DEBUGGER_STATE), (NAME, name)) => Some(name) case _ => None } } val DEBUGGER_OUTPUT = "debugger_output" object Debugger_Output { def unapply(props: Properties.T): Option[String] = props match { case List((FUNCTION, DEBUGGER_OUTPUT), (NAME, name)) => Some(name) case _ => None } } /* simplifier trace */ val SIMP_TRACE_PANEL = "simp_trace_panel" val SIMP_TRACE_LOG = "simp_trace_log" val SIMP_TRACE_STEP = "simp_trace_step" val SIMP_TRACE_RECURSE = "simp_trace_recurse" val SIMP_TRACE_HINT = "simp_trace_hint" val SIMP_TRACE_IGNORE = "simp_trace_ignore" val SIMP_TRACE_CANCEL = "simp_trace_cancel" object Simp_Trace_Cancel { def unapply(props: Properties.T): Option[Long] = props match { case (FUNCTION, SIMP_TRACE_CANCEL) :: Serial(i) => Some(i) case _ => None } } /* XML data representation */ def encode: XML.Encode.T[Markup] = { (markup: Markup) => import XML.Encode._ pair(string, properties)((markup.name, markup.properties)) } def decode: XML.Decode.T[Markup] = { (body: XML.Body) => import XML.Decode._ val (name, props) = pair(string, properties)(body) Markup(name, props) } } sealed case class Markup(name: String, properties: Properties.T) { def is_empty: Boolean = name.isEmpty def position_properties: Position.T = properties.filter(Markup.position_property) def markup(s: String): String = YXML.string_of_tree(XML.Elem(this, List(XML.Text(s)))) def update_properties(more_props: Properties.T): Markup = if (more_props.isEmpty) this else Markup(name, more_props.foldRight(properties) { case (p, ps) => Properties.put(ps, p) }) def + (entry: Properties.Entry): Markup = Markup(name, Properties.put(properties, entry)) }