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