| author | wenzelm | 
| Fri, 24 Sep 2021 22:23:26 +0200 | |
| changeset 74362 | 0135a0c77b64 | 
| parent 73340 | 0ffcad1f6130 | 
| child 74657 | 9fcf80ceb863 | 
| permissions | -rw-r--r-- | 
| 33984 | 1 | /* Title: Pure/Thy/html.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 50707 
5b2bf7611662
maintain session index on Scala side, for more determistic results;
 wenzelm parents: 
50201diff
changeset | 4 | HTML presentation elements. | 
| 33984 | 5 | */ | 
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 55618 | 9 | |
| 33984 | 10 | object HTML | 
| 11 | {
 | |
| 73207 | 12 | /* attributes */ | 
| 13 | ||
| 14 | class Attribute(val name: String, value: String) | |
| 15 |   {
 | |
| 16 | def xml: XML.Attribute = name -> value | |
| 17 | def apply(elem: XML.Elem): XML.Elem = elem + xml | |
| 18 | } | |
| 19 | ||
| 20 |   def id(s: String): Attribute = new Attribute("id", s)
 | |
| 21 |   def class_(name: String): Attribute = new Attribute("class", name)
 | |
| 22 | ||
| 23 |   def width(w: Int): Attribute = new Attribute("width", w.toString)
 | |
| 24 |   def height(h: Int): Attribute = new Attribute("height", h.toString)
 | |
| 25 | def size(w: Int, h: Int)(elem: XML.Elem): XML.Elem = width(w)(height(h)(elem)) | |
| 26 | ||
| 27 | ||
| 28 | /* structured markup operators */ | |
| 29 | ||
| 30 | def text(txt: String): XML.Body = if (txt.isEmpty) Nil else List(XML.Text(txt)) | |
| 31 |   val break: XML.Body = List(XML.elem("br"))
 | |
| 32 |   val nl: XML.Body = List(XML.Text("\n"))
 | |
| 33 | ||
| 34 | class Operator(val name: String) | |
| 35 |   {
 | |
| 36 | def apply(body: XML.Body): XML.Elem = XML.elem(name, body) | |
| 37 | def apply(att: Attribute, body: XML.Body): XML.Elem = att(apply(body)) | |
| 38 | def apply(c: String, body: XML.Body): XML.Elem = apply(class_(c), body) | |
| 39 | } | |
| 40 | ||
| 41 | class Heading(name: String) extends Operator(name) | |
| 42 |   {
 | |
| 43 | def apply(txt: String): XML.Elem = super.apply(text(txt)) | |
| 44 | def apply(att: Attribute, txt: String): XML.Elem = super.apply(att, text(txt)) | |
| 45 | def apply(c: String, txt: String): XML.Elem = super.apply(c, text(txt)) | |
| 46 | } | |
| 47 | ||
| 48 |   val div = new Operator("div")
 | |
| 49 |   val span = new Operator("span")
 | |
| 50 |   val pre = new Operator("pre")
 | |
| 51 |   val par = new Operator("p")
 | |
| 52 |   val sub = new Operator("sub")
 | |
| 53 |   val sup = new Operator("sup")
 | |
| 54 |   val emph = new Operator("em")
 | |
| 55 |   val bold = new Operator("b")
 | |
| 56 |   val code = new Operator("code")
 | |
| 57 |   val item = new Operator("li")
 | |
| 58 |   val list = new Operator("ul")
 | |
| 59 |   val enum = new Operator("ol")
 | |
| 60 |   val descr = new Operator("dl")
 | |
| 61 |   val dt = new Operator("dt")
 | |
| 62 |   val dd = new Operator("dd")
 | |
| 63 | ||
| 64 |   val title = new Heading("title")
 | |
| 65 |   val chapter = new Heading("h1")
 | |
| 66 |   val section = new Heading("h2")
 | |
| 67 |   val subsection = new Heading("h3")
 | |
| 68 |   val subsubsection = new Heading("h4")
 | |
| 69 |   val paragraph = new Heading("h5")
 | |
| 70 |   val subparagraph = new Heading("h6")
 | |
| 71 | ||
| 72 | def itemize(items: List[XML.Body]): XML.Elem = list(items.map(item(_))) | |
| 73 | def enumerate(items: List[XML.Body]): XML.Elem = enum(items.map(item(_))) | |
| 74 | def description(items: List[(XML.Body, XML.Body)]): XML.Elem = | |
| 75 |     descr(items.flatMap({ case (x, y) => List(dt(x), dd(y)) }))
 | |
| 76 | ||
| 77 | def link(href: String, body: XML.Body): XML.Elem = | |
| 78 |     XML.Elem(Markup("a", List("href" -> href)), if (body.isEmpty) text(href) else body)
 | |
| 79 | ||
| 80 | def link(path: Path, body: XML.Body): XML.Elem = link(path.implode, body) | |
| 81 | ||
| 82 | def image(src: String, alt: String = ""): XML.Elem = | |
| 83 |     XML.Elem(Markup("img", List("src" -> src) ::: proper_string(alt).map("alt" -> _).toList), Nil)
 | |
| 84 | ||
| 85 |   def source(body: XML.Body): XML.Elem = pre("source", body)
 | |
| 86 | def source(src: String): XML.Elem = source(text(src)) | |
| 87 | ||
| 88 |   def style(s: String): XML.Elem = XML.elem("style", text(s))
 | |
| 89 | def style_file(href: String): XML.Elem = | |
| 90 |     XML.Elem(Markup("link", List("rel" -> "stylesheet", "type" -> "text/css", "href" -> href)), Nil)
 | |
| 91 | def style_file(path: Path): XML.Elem = style_file(Url.print_file(path.file)) | |
| 92 | ||
| 93 |   def script(s: String): XML.Elem = XML.elem("script", text(s))
 | |
| 94 |   def script_file(href: String): XML.Elem = XML.Elem(Markup("script", List("src" -> href)), Nil)
 | |
| 95 | def script_file(path: Path): XML.Elem = script_file(Url.print_file(path.file)) | |
| 96 | ||
| 97 | ||
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 98 | /* output text with control symbols */ | 
| 59063 | 99 | |
| 73208 | 100 | private val control: Map[Symbol.Symbol, Operator] = | 
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 101 | Map( | 
| 73206 | 102 | Symbol.sub -> sub, Symbol.sub_decoded -> sub, | 
| 103 | Symbol.sup -> sup, Symbol.sup_decoded -> sup, | |
| 104 | Symbol.bold -> bold, Symbol.bold_decoded -> bold) | |
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 105 | |
| 73208 | 106 | private val control_block_begin: Map[Symbol.Symbol, Operator] = | 
| 62104 
fb73c0d7bb37
clarified symbol insertion, depending on buffer encoding;
 wenzelm parents: 
60215diff
changeset | 107 | Map( | 
| 73206 | 108 | Symbol.bsub -> sub, Symbol.bsub_decoded -> sub, | 
| 109 | Symbol.bsup -> sup, Symbol.bsup_decoded -> sup) | |
| 110 | ||
| 73208 | 111 | private val control_block_end: Map[Symbol.Symbol, Operator] = | 
| 73206 | 112 | Map( | 
| 113 | Symbol.esub -> sub, Symbol.esub_decoded -> sub, | |
| 114 | Symbol.esup -> sup, Symbol.esup_decoded -> sup) | |
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 115 | |
| 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 116 | def is_control(sym: Symbol.Symbol): Boolean = control.isDefinedAt(sym) | 
| 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 117 | |
| 73209 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 118 | def is_control_block_begin(sym: Symbol.Symbol): Boolean = control_block_begin.isDefinedAt(sym) | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 119 | def is_control_block_end(sym: Symbol.Symbol): Boolean = control_block_end.isDefinedAt(sym) | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 120 | def is_control_block_pair(bg: Symbol.Symbol, en: Symbol.Symbol): Boolean = | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 121 |   {
 | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 122 | val bg_decoded = Symbol.decode(bg) | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 123 | val en_decoded = Symbol.decode(en) | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 124 | bg_decoded == Symbol.bsub_decoded && en_decoded == Symbol.esub_decoded || | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 125 | bg_decoded == Symbol.bsup_decoded && en_decoded == Symbol.esup_decoded | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 126 | } | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 127 | |
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 128 | def check_control_blocks(body: XML.Body): Boolean = | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 129 |   {
 | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 130 | var ok = true | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 131 | var open = List.empty[Symbol.Symbol] | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 132 |     for { XML.Text(text) <- body; sym <- Symbol.iterator(text) } {
 | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 133 | if (is_control_block_begin(sym)) open ::= sym | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 134 |       else if (is_control_block_end(sym)) {
 | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 135 |         open match {
 | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 136 | case bg :: rest if is_control_block_pair(bg, sym) => open = rest | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 137 | case _ => ok = false | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 138 | } | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 139 | } | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 140 | } | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 141 | ok && open.isEmpty | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 142 | } | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 143 | |
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 144 | def output(s: StringBuilder, text: String, | 
| 73340 | 145 | control_blocks: Boolean, hidden: Boolean, permissive: Boolean): Unit = | 
| 66018 
9ce3720976dc
permissive output of XML.Text, e.g. relevant for embedded <style>;
 wenzelm parents: 
66006diff
changeset | 146 |   {
 | 
| 
9ce3720976dc
permissive output of XML.Text, e.g. relevant for embedded <style>;
 wenzelm parents: 
66006diff
changeset | 147 | def output_string(str: String): Unit = | 
| 73204 
aa3d4cf7825a
clarified signature: no symbol markup within XML attributes;
 wenzelm parents: 
73203diff
changeset | 148 | XML.output_string(s, str, permissive = permissive) | 
| 66018 
9ce3720976dc
permissive output of XML.Text, e.g. relevant for embedded <style>;
 wenzelm parents: 
66006diff
changeset | 149 | |
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 150 | def output_hidden(body: => Unit): Unit = | 
| 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 151 |       if (hidden) { s ++= "<span class=\"hidden\">"; body; s ++= "</span>" }
 | 
| 37200 
0f3edc64356a
added HTML.encode (in Scala), similar to HTML.output in ML;
 wenzelm parents: 
36016diff
changeset | 152 | |
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 153 | def output_symbol(sym: Symbol.Symbol): Unit = | 
| 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 154 |       if (sym != "") {
 | 
| 73206 | 155 |         control_block_begin.get(sym) match {
 | 
| 73209 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 156 | case Some(op) if control_blocks => | 
| 73206 | 157 | output_hidden(output_string(sym)) | 
| 158 | XML.output_elem(s, Markup(op.name, Nil)) | |
| 73209 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 159 | case _ => | 
| 73206 | 160 |             control_block_end.get(sym) match {
 | 
| 73209 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 161 | case Some(op) if control_blocks => | 
| 73206 | 162 | XML.output_elem_end(s, op.name) | 
| 163 | output_hidden(output_string(sym)) | |
| 73209 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 164 | case _ => | 
| 73206 | 165 |                 if (hidden && Symbol.is_control_encoded(sym)) {
 | 
| 166 | output_hidden(output_string(Symbol.control_prefix)) | |
| 167 | s ++= "<span class=\"control\">" | |
| 168 | output_string(Symbol.control_name(sym).get) | |
| 169 | s ++= "</span>" | |
| 170 | output_hidden(output_string(Symbol.control_suffix)) | |
| 171 | } | |
| 172 | else output_string(sym) | |
| 67255 
f1f983484878
HTML rendering of \<^control> as in Isabelle/jEdit;
 wenzelm parents: 
66212diff
changeset | 173 | } | 
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 174 | } | 
| 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 175 | } | 
| 59063 | 176 | |
| 62104 
fb73c0d7bb37
clarified symbol insertion, depending on buffer encoding;
 wenzelm parents: 
60215diff
changeset | 177 | var ctrl = "" | 
| 59063 | 178 |     for (sym <- Symbol.iterator(text)) {
 | 
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 179 |       if (is_control(sym)) { output_symbol(ctrl); ctrl = sym }
 | 
| 59063 | 180 |       else {
 | 
| 62104 
fb73c0d7bb37
clarified symbol insertion, depending on buffer encoding;
 wenzelm parents: 
60215diff
changeset | 181 |         control.get(ctrl) match {
 | 
| 73206 | 182 | case Some(op) if Symbol.is_controllable(sym) => | 
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 183 | output_hidden(output_symbol(ctrl)) | 
| 73206 | 184 | XML.output_elem(s, Markup(op.name, Nil)) | 
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 185 | output_symbol(sym) | 
| 73206 | 186 | XML.output_elem_end(s, op.name) | 
| 59063 | 187 | case _ => | 
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 188 | output_symbol(ctrl) | 
| 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 189 | output_symbol(sym) | 
| 59063 | 190 | } | 
| 62104 
fb73c0d7bb37
clarified symbol insertion, depending on buffer encoding;
 wenzelm parents: 
60215diff
changeset | 191 | ctrl = "" | 
| 59063 | 192 | } | 
| 37200 
0f3edc64356a
added HTML.encode (in Scala), similar to HTML.output in ML;
 wenzelm parents: 
36016diff
changeset | 193 | } | 
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 194 | output_symbol(ctrl) | 
| 64355 | 195 | } | 
| 59063 | 196 | |
| 66018 
9ce3720976dc
permissive output of XML.Text, e.g. relevant for embedded <style>;
 wenzelm parents: 
66006diff
changeset | 197 | def output(text: String): String = | 
| 73209 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 198 |   {
 | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 199 | val control_blocks = check_control_blocks(List(XML.Text(text))) | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 200 | Library.make_string(output(_, text, | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 201 | control_blocks = control_blocks, hidden = false, permissive = true)) | 
| 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 202 | } | 
| 64355 | 203 | |
| 204 | ||
| 205 | /* output XML as HTML */ | |
| 206 | ||
| 65834 
67a6e0f166c2
extra space only for some structual elements, but not <a>, <b>, <em> etc. (amending 8a0fe5469ba0);
 wenzelm parents: 
65773diff
changeset | 207 | private val structural_elements = | 
| 
67a6e0f166c2
extra space only for some structual elements, but not <a>, <b>, <em> etc. (amending 8a0fe5469ba0);
 wenzelm parents: 
65773diff
changeset | 208 |     Set("head", "body", "meta", "div", "pre", "p", "title", "h1", "h2", "h3", "h4", "h5", "h6",
 | 
| 
67a6e0f166c2
extra space only for some structual elements, but not <a>, <b>, <em> etc. (amending 8a0fe5469ba0);
 wenzelm parents: 
65773diff
changeset | 209 | "ul", "ol", "dl", "li", "dt", "dd") | 
| 
67a6e0f166c2
extra space only for some structual elements, but not <a>, <b>, <em> etc. (amending 8a0fe5469ba0);
 wenzelm parents: 
65773diff
changeset | 210 | |
| 73340 | 211 | def output(s: StringBuilder, xml: XML.Body, hidden: Boolean, structural: Boolean): Unit = | 
| 64355 | 212 |   {
 | 
| 73205 | 213 | def output_body(body: XML.Body): Unit = | 
| 214 |     {
 | |
| 73209 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 215 | val control_blocks = check_control_blocks(body) | 
| 73205 | 216 |       body foreach {
 | 
| 64355 | 217 | case XML.Elem(markup, Nil) => | 
| 73204 
aa3d4cf7825a
clarified signature: no symbol markup within XML attributes;
 wenzelm parents: 
73203diff
changeset | 218 | XML.output_elem(s, markup, end = true) | 
| 64355 | 219 | case XML.Elem(markup, ts) => | 
| 67337 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 220 | if (structural && structural_elements(markup.name)) s += '\n' | 
| 73204 
aa3d4cf7825a
clarified signature: no symbol markup within XML attributes;
 wenzelm parents: 
73203diff
changeset | 221 | XML.output_elem(s, markup) | 
| 73205 | 222 | output_body(ts) | 
| 73204 
aa3d4cf7825a
clarified signature: no symbol markup within XML attributes;
 wenzelm parents: 
73203diff
changeset | 223 | XML.output_elem_end(s, markup.name) | 
| 67337 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 224 | if (structural && structural_elements(markup.name)) s += '\n' | 
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 225 | case XML.Text(txt) => | 
| 73209 
de16d797adbe
more robust HTML output: non-balanced bsub/esub are shown verbatim, e.g. within mixfix declarations (due to extra markup);
 wenzelm parents: 
73208diff
changeset | 226 | output(s, txt, control_blocks = control_blocks, hidden = hidden, permissive = true) | 
| 64355 | 227 | } | 
| 73205 | 228 | } | 
| 229 | output_body(xml) | |
| 37200 
0f3edc64356a
added HTML.encode (in Scala), similar to HTML.output in ML;
 wenzelm parents: 
36016diff
changeset | 230 | } | 
| 
0f3edc64356a
added HTML.encode (in Scala), similar to HTML.output in ML;
 wenzelm parents: 
36016diff
changeset | 231 | |
| 67337 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 232 | def output(body: XML.Body, hidden: Boolean, structural: Boolean): String = | 
| 73202 | 233 | Library.make_string(output(_, body, hidden, structural)) | 
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 234 | |
| 67337 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 235 | def output(tree: XML.Tree, hidden: Boolean, structural: Boolean): String = | 
| 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 236 | output(List(tree), hidden, structural) | 
| 64355 | 237 | |
| 238 | ||
| 65892 | 239 | /* messages */ | 
| 240 | ||
| 65939 | 241 | // background | 
| 65993 | 242 |   val writeln_message: Attribute = class_("writeln_message")
 | 
| 243 |   val warning_message: Attribute = class_("warning_message")
 | |
| 244 |   val error_message: Attribute = class_("error_message")
 | |
| 65892 | 245 | |
| 65939 | 246 | // underline | 
| 65993 | 247 |   val writeln: Attribute = class_("writeln")
 | 
| 248 |   val warning: Attribute = class_("warning")
 | |
| 249 |   val error: Attribute = class_("error")
 | |
| 65892 | 250 | |
| 65944 | 251 | |
| 252 | /* tooltips */ | |
| 253 | ||
| 254 | def tooltip(item: XML.Body, tip: XML.Body): XML.Elem = | |
| 65992 | 255 |     span(item ::: List(div("tooltip", tip)))
 | 
| 65944 | 256 | |
| 257 | def tooltip_errors(item: XML.Body, msgs: List[XML.Body]): XML.Elem = | |
| 258 | HTML.error(tooltip(item, msgs.map(msg => error_message(pre(msg))))) | |
| 65892 | 259 | |
| 260 | ||
| 66196 | 261 | /* GUI elements */ | 
| 262 | ||
| 263 | object GUI | |
| 264 |   {
 | |
| 265 | private def optional_value(text: String): XML.Attributes = | |
| 266 | proper_string(text).map(a => "value" -> a).toList | |
| 267 | ||
| 268 | private def optional_name(name: String): XML.Attributes = | |
| 269 | proper_string(name).map(a => "name" -> a).toList | |
| 270 | ||
| 271 | private def optional_title(tooltip: String): XML.Attributes = | |
| 272 | proper_string(tooltip).map(a => "title" -> a).toList | |
| 273 | ||
| 274 | private def optional_submit(submit: Boolean): XML.Attributes = | |
| 275 |       if (submit) List("onChange" -> "this.form.submit()") else Nil
 | |
| 276 | ||
| 277 | private def optional_checked(selected: Boolean): XML.Attributes = | |
| 278 |       if (selected) List("checked" -> "") else Nil
 | |
| 279 | ||
| 280 | private def optional_action(action: String): XML.Attributes = | |
| 281 | proper_string(action).map(a => "action" -> a).toList | |
| 282 | ||
| 66210 | 283 | private def optional_onclick(script: String): XML.Attributes = | 
| 284 | proper_string(script).map(a => "onclick" -> a).toList | |
| 285 | ||
| 286 | private def optional_onchange(script: String): XML.Attributes = | |
| 287 | proper_string(script).map(a => "onchange" -> a).toList | |
| 288 | ||
| 289 | def button(body: XML.Body, name: String = "", tooltip: String = "", submit: Boolean = false, | |
| 290 | script: String = ""): XML.Elem = | |
| 66196 | 291 | XML.Elem( | 
| 292 |         Markup("button",
 | |
| 293 |           List("type" -> (if (submit) "submit" else "button"), "value" -> "true") :::
 | |
| 66210 | 294 | optional_name(name) ::: optional_title(tooltip) ::: optional_onclick(script)), body) | 
| 66196 | 295 | |
| 296 | def checkbox(body: XML.Body, name: String = "", tooltip: String = "", submit: Boolean = false, | |
| 66210 | 297 | selected: Boolean = false, script: String = ""): XML.Elem = | 
| 66196 | 298 |       XML.elem("label",
 | 
| 299 | XML.elem( | |
| 300 |           Markup("input",
 | |
| 301 |             List("type" -> "checkbox", "value" -> "true") ::: optional_name(name) :::
 | |
| 302 | optional_title(tooltip) ::: optional_submit(submit) ::: | |
| 66210 | 303 | optional_checked(selected) ::: optional_onchange(script))) :: body) | 
| 66196 | 304 | |
| 305 | def text_field(columns: Int = 0, text: String = "", name: String = "", tooltip: String = "", | |
| 66210 | 306 | submit: Boolean = false, script: String = ""): XML.Elem = | 
| 66196 | 307 |       XML.elem(Markup("input",
 | 
| 308 |         List("type" -> "text") :::
 | |
| 309 |           (if (columns > 0) List("size" -> columns.toString) else Nil) :::
 | |
| 66210 | 310 | optional_value(text) ::: optional_name(name) ::: optional_title(tooltip) ::: | 
| 311 | optional_submit(submit) ::: optional_onchange(script))) | |
| 66196 | 312 | |
| 313 | def parameter(text: String = "", name: String = ""): XML.Elem = | |
| 314 | XML.elem( | |
| 315 |         Markup("input", List("type" -> "hidden") ::: optional_value(text) ::: optional_name(name)))
 | |
| 316 | ||
| 66209 | 317 | def form(body: XML.Body, name: String = "", action: String = "", http_post: Boolean = false) | 
| 318 | : XML.Elem = | |
| 319 | XML.Elem( | |
| 320 |         Markup("form", optional_name(name) ::: optional_action(action) :::
 | |
| 321 |           (if (http_post) List("method" -> "post") else Nil)), body)
 | |
| 66196 | 322 | } | 
| 323 | ||
| 324 | ||
| 66200 | 325 | /* GUI layout */ | 
| 326 | ||
| 327 | object Wrap_Panel | |
| 328 |   {
 | |
| 329 | object Alignment extends Enumeration | |
| 330 |     {
 | |
| 331 | val left, right, center = Value | |
| 332 | } | |
| 333 | ||
| 334 | def apply(contents: List[XML.Elem], name: String = "", action: String = "", | |
| 66206 | 335 | alignment: Alignment.Value = Alignment.right): XML.Elem = | 
| 66200 | 336 |     {
 | 
| 337 |       val body = Library.separate(XML.Text(" "), contents)
 | |
| 338 |       GUI.form(List(div(body) + ("style" -> ("text-align: " + alignment))),
 | |
| 339 | name = name, action = action) | |
| 340 | } | |
| 341 | } | |
| 342 | ||
| 343 | ||
| 51399 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 wenzelm parents: 
50707diff
changeset | 344 | /* document */ | 
| 
6ac3c29a300e
discontinued "isabelle usedir" option -r (reset session path);
 wenzelm parents: 
50707diff
changeset | 345 | |
| 64359 | 346 | val header: String = | 
| 69804 | 347 | XML.header + | 
| 348 | """<!DOCTYPE html PUBLIC "-//W3C//DTD XHTML 1.0 Transitional//EN" "http://www.w3.org/TR/xhtml1/DTD/xhtml1-transitional.dtd"> | |
| 64359 | 349 | <html xmlns="http://www.w3.org/1999/xhtml">""" | 
| 350 | ||
| 73129 | 351 | val footer: String = """</html>""" | 
| 352 | ||
| 64359 | 353 | val head_meta: XML.Elem = | 
| 354 |     XML.Elem(Markup("meta",
 | |
| 355 |       List("http-equiv" -> "Content-Type", "content" -> "text/html; charset=utf-8")), Nil)
 | |
| 356 | ||
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 357 | def output_document(head: XML.Body, body: XML.Body, | 
| 67337 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 358 | css: String = "isabelle.css", | 
| 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 359 | hidden: Boolean = true, | 
| 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 360 | structural: Boolean = true): String = | 
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 361 |   {
 | 
| 64359 | 362 | cat_lines( | 
| 73129 | 363 | List( | 
| 364 | header, | |
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 365 | output( | 
| 66000 | 366 |           XML.elem("head", head_meta :: (if (css == "") Nil else List(style_file(css))) ::: head),
 | 
| 67337 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 367 | hidden = hidden, structural = structural), | 
| 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 368 |         output(XML.elem("body", body),
 | 
| 73129 | 369 | hidden = hidden, structural = structural), | 
| 370 | footer)) | |
| 65997 
e3dc9ea67a62
output control symbols like ML version, with optionally hidden source;
 wenzelm parents: 
65993diff
changeset | 371 | } | 
| 64359 | 372 | |
| 65838 | 373 | |
| 66000 | 374 | /* fonts */ | 
| 375 | ||
| 376 | def fonts_url(): String => String = | |
| 69374 | 377 | (for (entry <- Isabelle_Fonts.fonts(hidden = true)) | 
| 70743 
342b0a1fc86d
proper file name instead of font name (amending dc9a39c3f75d);
 wenzelm parents: 
69804diff
changeset | 378 | yield (entry.path.file_name -> Url.print_file(entry.path.file))).toMap | 
| 66000 | 379 | |
| 380 | def fonts_dir(prefix: String)(ttf_name: String): String = | |
| 381 | prefix + "/" + ttf_name | |
| 382 | ||
| 383 | def fonts_css(make_url: String => String = fonts_url()): String = | |
| 384 |   {
 | |
| 69360 | 385 | def font_face(entry: Isabelle_Fonts.Entry): String = | 
| 66000 | 386 | cat_lines( | 
| 387 | List( | |
| 388 |           "@font-face {",
 | |
| 69360 | 389 | " font-family: '" + entry.family + "';", | 
| 69366 | 390 |           "  src: url('" + make_url(entry.path.file_name) + "') format('truetype');") :::
 | 
| 69360 | 391 |         (if (entry.is_bold) List("  font-weight: bold;") else Nil) :::
 | 
| 392 |         (if (entry.is_italic) List("  font-style: italic;") else Nil) :::
 | |
| 66000 | 393 |         List("}"))
 | 
| 394 | ||
| 71601 | 395 |     ("/* Isabelle fonts */" :: Isabelle_Fonts.fonts(hidden = true).map(font_face))
 | 
| 69362 | 396 |       .mkString("", "\n\n", "\n")
 | 
| 66000 | 397 | } | 
| 398 | ||
| 399 | ||
| 65838 | 400 | /* document directory */ | 
| 401 | ||
| 65998 | 402 |   def isabelle_css: Path = Path.explode("~~/etc/isabelle.css")
 | 
| 403 | ||
| 73340 | 404 |   def write_isabelle_css(dir: Path, make_url: String => String = fonts_dir("fonts")): Unit =
 | 
| 66000 | 405 |   {
 | 
| 406 | File.write(dir + isabelle_css.base, | |
| 69362 | 407 | fonts_css(make_url) + "\n\n" + File.read(isabelle_css)) | 
| 66000 | 408 | } | 
| 409 | ||
| 72376 | 410 | def init_dir(dir: Path): Unit = | 
| 411 | write_isabelle_css(Isabelle_System.make_directory(dir)) | |
| 65836 | 412 | |
| 65838 | 413 | def write_document(dir: Path, name: String, head: XML.Body, body: XML.Body, | 
| 69366 | 414 | css: String = isabelle_css.file_name, | 
| 67337 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 415 | hidden: Boolean = true, | 
| 73340 | 416 | structural: Boolean = true): Unit = | 
| 65838 | 417 |   {
 | 
| 418 | init_dir(dir) | |
| 67337 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 419 | File.write(dir + Path.basic(name), | 
| 
4254cfd15b00
more tight HTML output: avoid extra lines within <pre>;
 wenzelm parents: 
67336diff
changeset | 420 | output_document(head, body, css = css, hidden = hidden, structural = structural)) | 
| 65838 | 421 | } | 
| 33984 | 422 | } |