| author | haftmann | 
| Fri, 30 Mar 2012 09:04:29 +0200 | |
| changeset 47202 | 69cee87927f0 | 
| parent 46866 | b190913c3c41 | 
| child 48670 | 206144b13849 | 
| permissions | -rw-r--r-- | 
| 24584 | 1 | (* Title: Pure/Thy/html.ML | 
| 9415 | 2 | Author: Markus Wenzel and Stefan Berghofer, TU Muenchen | 
| 6324 | 3 | |
| 9415 | 4 | HTML presentation elements. | 
| 6324 | 5 | *) | 
| 6 | ||
| 7 | signature HTML = | |
| 8 | sig | |
| 23724 | 9 |   val html_mode: ('a -> 'b) -> 'a -> 'b
 | 
| 23622 | 10 | type text = string | 
| 6324 | 11 | val plain: string -> text | 
| 12 | val name: string -> text | |
| 13 | val keyword: string -> text | |
| 23622 | 14 | val command: string -> text | 
| 6324 | 15 | val file_name: string -> text | 
| 6649 | 16 | val file_path: Url.T -> text | 
| 6324 | 17 | val href_name: string -> text -> text | 
| 6649 | 18 | val href_path: Url.T -> text -> text | 
| 19 | val href_opt_path: Url.T option -> text -> text | |
| 6376 | 20 | val para: text -> text | 
| 6324 | 21 | val preform: text -> text | 
| 22 | val verbatim: string -> text | |
| 23724 | 23 | val begin_document: string -> text | 
| 24 | val end_document: text | |
| 17080 | 25 | val begin_index: Url.T * string -> Url.T * string -> (Url.T * string) list -> Url.T -> text | 
| 9415 | 26 | val applet_pages: string -> Url.T * string -> (string * string) list | 
| 6649 | 27 | val theory_entry: Url.T * string -> text | 
| 28 | val session_entries: (Url.T * string) list -> text | |
| 27861 
911bf8e58c4c
removed obsolete verbatim_source, results, chapter, section etc.;
 wenzelm parents: 
27832diff
changeset | 29 | val theory_source: text -> text | 
| 6649 | 30 | val begin_theory: Url.T * string -> string -> (Url.T option * string) list -> | 
| 27490 | 31 | (Url.T * Url.T * bool) list -> text -> text | 
| 37940 
4857eab31298
generic external source files -- nothing special about ML here;
 wenzelm parents: 
37146diff
changeset | 32 | val external_file: Url.T -> string -> text | 
| 6324 | 33 | end; | 
| 34 | ||
| 35 | structure HTML: HTML = | |
| 36 | struct | |
| 37 | ||
| 38 | ||
| 39 | (** HTML print modes **) | |
| 40 | ||
| 41 | (* mode *) | |
| 42 | ||
| 43 | val htmlN = "HTML"; | |
| 37146 
f652333bbf8e
renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
 wenzelm parents: 
34003diff
changeset | 44 | fun html_mode f x = Print_Mode.with_modes [htmlN] f x; | 
| 6324 | 45 | |
| 46 | ||
| 33985 | 47 | (* common markup *) | 
| 48 | ||
| 49 | fun span class = ("<span class=" ^ quote (XML.text class) ^ ">", "</span>");
 | |
| 50 | ||
| 46866 
b190913c3c41
simplified span class in conformance to Scala version;
 wenzelm parents: 
45666diff
changeset | 51 | val _ = Markup.add_mode htmlN (span o fst); | 
| 33985 | 52 | |
| 53 | ||
| 6324 | 54 | (* symbol output *) | 
| 55 | ||
| 6338 | 56 | local | 
| 45666 | 57 | val hidden = span Isabelle_Markup.hiddenN |-> enclose; | 
| 33985 | 58 | |
| 43592 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 59 | (* FIXME proper unicode -- produced on Scala side *) | 
| 16196 | 60 | val html_syms = Symtab.make | 
| 34003 
610e41138486
output_syms: permissive treatment of control symbols, cf. Scala version;
 wenzelm parents: 
33986diff
changeset | 61 |    [("", (0, "")),
 | 
| 
610e41138486
output_syms: permissive treatment of control symbols, cf. Scala version;
 wenzelm parents: 
33986diff
changeset | 62 |     ("\n", (0, "<br/>")),
 | 
| 33986 
041dc6d8d344
output linefeed as </br> -- workaround problem with <pre> in Lobo Browser 0.98.4;
 wenzelm parents: 
33985diff
changeset | 63 |     ("'", (1, "'")),
 | 
| 33983 
cfbf1ff6170d
output "'" as "'" which is a bit more portable ("'" is defined in XML/XHTML, but not in old-style HTML4);
 wenzelm parents: 
33222diff
changeset | 64 |     ("\\<spacespace>", (2, "  ")),
 | 
| 23622 | 65 |     ("\\<exclamdown>", (1, "¡")),
 | 
| 66 |     ("\\<cent>", (1, "¢")),
 | |
| 67 |     ("\\<pounds>", (1, "£")),
 | |
| 68 |     ("\\<currency>", (1, "¤")),
 | |
| 69 |     ("\\<yen>", (1, "¥")),
 | |
| 70 |     ("\\<bar>", (1, "¦")),
 | |
| 71 |     ("\\<section>", (1, "§")),
 | |
| 72 |     ("\\<dieresis>", (1, "¨")),
 | |
| 73 |     ("\\<copyright>", (1, "©")),
 | |
| 74 |     ("\\<ordfeminine>", (1, "ª")),
 | |
| 75 |     ("\\<guillemotleft>", (1, "«")),
 | |
| 76 |     ("\\<not>", (1, "¬")),
 | |
| 77 |     ("\\<hyphen>", (1, "­")),
 | |
| 78 |     ("\\<registered>", (1, "®")),
 | |
| 79 |     ("\\<inverse>", (1, "¯")),
 | |
| 80 |     ("\\<degree>", (1, "°")),
 | |
| 81 |     ("\\<plusminus>", (1, "±")),
 | |
| 82 |     ("\\<twosuperior>", (1, "²")),
 | |
| 83 |     ("\\<threesuperior>", (1, "³")),
 | |
| 84 |     ("\\<acute>", (1, "´")),
 | |
| 85 |     ("\\<paragraph>", (1, "¶")),
 | |
| 86 |     ("\\<cdot>", (1, "·")),
 | |
| 87 |     ("\\<cedilla>", (1, "¸")),
 | |
| 88 |     ("\\<onesuperior>", (1, "¹")),
 | |
| 89 |     ("\\<ordmasculine>", (1, "º")),
 | |
| 90 |     ("\\<guillemotright>", (1, "»")),
 | |
| 91 |     ("\\<onequarter>", (1, "¼")),
 | |
| 92 |     ("\\<onehalf>", (1, "½")),
 | |
| 93 |     ("\\<threequarters>", (1, "¾")),
 | |
| 94 |     ("\\<questiondown>", (1, "¿")),
 | |
| 95 |     ("\\<times>", (1, "×")),
 | |
| 96 |     ("\\<div>", (1, "÷")),
 | |
| 97 |     ("\\<circ>", (1, "o")),
 | |
| 98 |     ("\\<Alpha>", (1, "Α")),
 | |
| 99 |     ("\\<Beta>", (1, "Β")),
 | |
| 100 |     ("\\<Gamma>", (1, "Γ")),
 | |
| 101 |     ("\\<Delta>", (1, "Δ")),
 | |
| 102 |     ("\\<Epsilon>", (1, "Ε")),
 | |
| 103 |     ("\\<Zeta>", (1, "Ζ")),
 | |
| 104 |     ("\\<Eta>", (1, "Η")),
 | |
| 105 |     ("\\<Theta>", (1, "Θ")),
 | |
| 106 |     ("\\<Iota>", (1, "Ι")),
 | |
| 107 |     ("\\<Kappa>", (1, "Κ")),
 | |
| 108 |     ("\\<Lambda>", (1, "Λ")),
 | |
| 109 |     ("\\<Mu>", (1, "Μ")),
 | |
| 110 |     ("\\<Nu>", (1, "Ν")),
 | |
| 111 |     ("\\<Xi>", (1, "Ξ")),
 | |
| 112 |     ("\\<Omicron>", (1, "Ο")),
 | |
| 113 |     ("\\<Pi>", (1, "Π")),
 | |
| 114 |     ("\\<Rho>", (1, "Ρ")),
 | |
| 115 |     ("\\<Sigma>", (1, "Σ")),
 | |
| 116 |     ("\\<Tau>", (1, "Τ")),
 | |
| 117 |     ("\\<Upsilon>", (1, "Υ")),
 | |
| 118 |     ("\\<Phi>", (1, "Φ")),
 | |
| 119 |     ("\\<Chi>", (1, "Χ")),
 | |
| 120 |     ("\\<Psi>", (1, "Ψ")),
 | |
| 121 |     ("\\<Omega>", (1, "Ω")),
 | |
| 122 |     ("\\<alpha>", (1, "α")),
 | |
| 123 |     ("\\<beta>", (1, "β")),
 | |
| 124 |     ("\\<gamma>", (1, "γ")),
 | |
| 125 |     ("\\<delta>", (1, "δ")),
 | |
| 126 |     ("\\<epsilon>", (1, "ε")),
 | |
| 127 |     ("\\<zeta>", (1, "ζ")),
 | |
| 128 |     ("\\<eta>", (1, "η")),
 | |
| 129 |     ("\\<theta>", (1, "ϑ")),
 | |
| 130 |     ("\\<iota>", (1, "ι")),
 | |
| 131 |     ("\\<kappa>", (1, "κ")),
 | |
| 132 |     ("\\<lambda>", (1, "λ")),
 | |
| 133 |     ("\\<mu>", (1, "μ")),
 | |
| 134 |     ("\\<nu>", (1, "ν")),
 | |
| 135 |     ("\\<xi>", (1, "ξ")),
 | |
| 136 |     ("\\<omicron>", (1, "ο")),
 | |
| 137 |     ("\\<pi>", (1, "π")),
 | |
| 138 |     ("\\<rho>", (1, "ρ")),
 | |
| 139 |     ("\\<sigma>", (1, "σ")),
 | |
| 140 |     ("\\<tau>", (1, "τ")),
 | |
| 141 |     ("\\<upsilon>", (1, "υ")),
 | |
| 142 |     ("\\<phi>", (1, "φ")),
 | |
| 143 |     ("\\<chi>", (1, "χ")),
 | |
| 144 |     ("\\<psi>", (1, "ψ")),
 | |
| 145 |     ("\\<omega>", (1, "ω")),
 | |
| 146 |     ("\\<bullet>", (1, "•")),
 | |
| 147 |     ("\\<dots>", (1, "…")),
 | |
| 148 |     ("\\<Re>", (1, "ℜ")),
 | |
| 149 |     ("\\<Im>", (1, "ℑ")),
 | |
| 150 |     ("\\<wp>", (1, "℘")),
 | |
| 151 |     ("\\<forall>", (1, "∀")),
 | |
| 152 |     ("\\<partial>", (1, "∂")),
 | |
| 153 |     ("\\<exists>", (1, "∃")),
 | |
| 154 |     ("\\<emptyset>", (1, "∅")),
 | |
| 155 |     ("\\<nabla>", (1, "∇")),
 | |
| 156 |     ("\\<in>", (1, "∈")),
 | |
| 157 |     ("\\<notin>", (1, "∉")),
 | |
| 158 |     ("\\<Prod>", (1, "∏")),
 | |
| 159 |     ("\\<Sum>", (1, "∑")),
 | |
| 160 |     ("\\<star>", (1, "∗")),
 | |
| 161 |     ("\\<propto>", (1, "∝")),
 | |
| 162 |     ("\\<infinity>", (1, "∞")),
 | |
| 163 |     ("\\<angle>", (1, "∠")),
 | |
| 164 |     ("\\<and>", (1, "∧")),
 | |
| 165 |     ("\\<or>", (1, "∨")),
 | |
| 166 |     ("\\<inter>", (1, "∩")),
 | |
| 167 |     ("\\<union>", (1, "∪")),
 | |
| 168 |     ("\\<sim>", (1, "∼")),
 | |
| 169 |     ("\\<cong>", (1, "≅")),
 | |
| 170 |     ("\\<approx>", (1, "≈")),
 | |
| 171 |     ("\\<noteq>", (1, "≠")),
 | |
| 172 |     ("\\<equiv>", (1, "≡")),
 | |
| 173 |     ("\\<le>", (1, "≤")),
 | |
| 174 |     ("\\<ge>", (1, "≥")),
 | |
| 175 |     ("\\<subset>", (1, "⊂")),
 | |
| 176 |     ("\\<supset>", (1, "⊃")),
 | |
| 177 |     ("\\<subseteq>", (1, "⊆")),
 | |
| 178 |     ("\\<supseteq>", (1, "⊇")),
 | |
| 179 |     ("\\<oplus>", (1, "⊕")),
 | |
| 180 |     ("\\<otimes>", (1, "⊗")),
 | |
| 181 |     ("\\<bottom>", (1, "⊥")),
 | |
| 182 |     ("\\<lceil>", (1, "⌈")),
 | |
| 183 |     ("\\<rceil>", (1, "⌉")),
 | |
| 184 |     ("\\<lfloor>", (1, "⌊")),
 | |
| 185 |     ("\\<rfloor>", (1, "⌋")),
 | |
| 186 |     ("\\<langle>", (1, "⟨")),
 | |
| 187 |     ("\\<rangle>", (1, "⟩")),
 | |
| 188 |     ("\\<lozenge>", (1, "◊")),
 | |
| 189 |     ("\\<spadesuit>", (1, "♠")),
 | |
| 190 |     ("\\<clubsuit>", (1, "♣")),
 | |
| 191 |     ("\\<heartsuit>", (1, "♥")),
 | |
| 192 |     ("\\<diamondsuit>", (1, "♦")),
 | |
| 193 |     ("\\<lbrakk>", (2, "[|")),
 | |
| 194 |     ("\\<rbrakk>", (2, "|]")),
 | |
| 195 |     ("\\<Longrightarrow>", (3, "==>")),
 | |
| 196 |     ("\\<Rightarrow>", (2, "=>")),
 | |
| 197 |     ("\\<And>", (2, "!!")),
 | |
| 198 |     ("\\<Colon>", (2, "::")),
 | |
| 199 |     ("\\<lparr>", (2, "(|")),
 | |
| 200 |     ("\\<rparr>", (2, "|)),")),
 | |
| 201 |     ("\\<longleftrightarrow>", (3, "<->")),
 | |
| 202 |     ("\\<longrightarrow>", (3, "-->")),
 | |
| 203 |     ("\\<rightarrow>", (2, "->")),
 | |
| 43592 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 204 |     ("\\<^bsub>", (0, hidden "⇘" ^ "<sub>")),
 | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 205 |     ("\\<^esub>", (0, hidden "⇙" ^ "</sub>")),
 | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 206 |     ("\\<^bsup>", (0, hidden "⇗" ^ "<sup>")),
 | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 207 |     ("\\<^esup>", (0, hidden "⇖" ^ "</sup>"))];
 | 
| 17178 | 208 | |
| 16196 | 209 | fun output_sym s = | 
| 23622 | 210 | if Symbol.is_raw s then (1, Symbol.decode_raw s) | 
| 20742 | 211 | else | 
| 33983 
cfbf1ff6170d
output "'" as "'" which is a bit more portable ("'" is defined in XML/XHTML, but not in old-style HTML4);
 wenzelm parents: 
33222diff
changeset | 212 | (case Symtab.lookup html_syms s of | 
| 
cfbf1ff6170d
output "'" as "'" which is a bit more portable ("'" is defined in XML/XHTML, but not in old-style HTML4);
 wenzelm parents: 
33222diff
changeset | 213 | SOME x => x | 
| 26709 | 214 | | NONE => (size s, XML.text s)); | 
| 6324 | 215 | |
| 33985 | 216 | fun output_markup (bg, en) s1 s2 = | 
| 217 | let val (n, txt) = output_sym s2 | |
| 218 | in (n, hidden s1 ^ enclose bg en txt) end; | |
| 14534 | 219 | |
| 33985 | 220 |   val output_sub = output_markup ("<sub>", "</sub>");
 | 
| 221 |   val output_sup = output_markup ("<sup>", "</sup>");
 | |
| 222 | val output_bold = output_markup (span "bold"); | |
| 223 | ||
| 41482 | 224 | fun output_syms [] (result, width) = (implode (rev result), width) | 
| 225 | | output_syms (s1 :: rest) (result, width) = | |
| 226 | let | |
| 227 |           val (s2, ss) = (case rest of [] => ("", []) | s2 :: ss => (s2, ss));
 | |
| 228 | val ((w, s), r) = | |
| 43592 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 229 | if s1 = "\\<^sub>" then (output_sub "⇩" s2, ss) | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 230 | else if s1 = "\\<^isub>" then (output_sub "⇣" s2, ss) | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 231 | else if s1 = "\\<^sup>" then (output_sup "⇧" s2, ss) | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 232 | else if s1 = "\\<^isup>" then (output_sup "⇡" s2, ss) | 
| 
e67d104c0c50
HTML: render control symbols more like Isabelle/Scala/jEdit;
 wenzelm parents: 
43548diff
changeset | 233 | else if s1 = "\\<^bold>" then (output_bold "❙" s2, ss) | 
| 41482 | 234 | else (output_sym s1, rest); | 
| 235 | in output_syms r (s :: result, width + w) end; | |
| 6338 | 236 | in | 
| 6324 | 237 | |
| 41482 | 238 | fun output_width str = output_syms (Symbol.explode str) ([], 0); | 
| 6338 | 239 | val output = #1 o output_width; | 
| 240 | ||
| 26709 | 241 | val _ = Output.add_mode htmlN output_width Symbol.encode_raw; | 
| 242 | ||
| 6338 | 243 | end; | 
| 244 | ||
| 6324 | 245 | |
| 246 | ||
| 247 | (** HTML markup **) | |
| 248 | ||
| 249 | type text = string; | |
| 250 | ||
| 251 | ||
| 252 | (* atoms *) | |
| 253 | ||
| 254 | val plain = output; | |
| 19265 | 255 | val name = enclose "<span class=\"name\">" "</span>" o output; | 
| 256 | val keyword = enclose "<span class=\"keyword\">" "</span>" o output; | |
| 23622 | 257 | val command = enclose "<span class=\"command\">" "</span>" o output; | 
| 19265 | 258 | val file_name = enclose "<span class=\"filename\">" "</span>" o output; | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
20742diff
changeset | 259 | val file_path = file_name o Url.implode; | 
| 6324 | 260 | |
| 261 | ||
| 262 | (* misc *) | |
| 263 | ||
| 28840 | 264 | fun href_name s txt = "<a href=" ^ quote s ^ ">" ^ txt ^ "</a>"; | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
20742diff
changeset | 265 | fun href_path path txt = href_name (Url.implode path) txt; | 
| 6376 | 266 | |
| 15531 | 267 | fun href_opt_path NONE txt = txt | 
| 268 | | href_opt_path (SOME p) txt = href_path p txt; | |
| 6376 | 269 | |
| 12413 | 270 | fun para txt = "\n<p>" ^ txt ^ "</p>\n"; | 
| 27830 | 271 | fun preform txt = "<pre>" ^ txt ^ "</pre>"; | 
| 6324 | 272 | val verbatim = preform o output; | 
| 273 | ||
| 274 | ||
| 275 | (* document *) | |
| 276 | ||
| 277 | fun begin_document title = | |
| 40946 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 278 | "<?xml version=\"1.0\" encoding=\"utf-8\" ?>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 279 | \<!DOCTYPE html PUBLIC \"-//W3C//DTD XHTML 1.0 Transitional//EN\" \ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 280 | \\"http://www.w3.org/TR/xhtml1/DTD/xhtml1-transitional.dtd\">\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 281 | \<html xmlns=\"http://www.w3.org/1999/xhtml\">\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 282 | \<head>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 283 | \<meta http-equiv=\"Content-Type\" content=\"text/html; charset=utf-8\"/>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 284 |   \<title>" ^ plain (title ^ " (" ^ Distribution.version ^ ")") ^ "</title>\n\
 | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 285 | \<link media=\"all\" rel=\"stylesheet\" type=\"text/css\" href=\"isabelle.css\"/>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 286 | \</head>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 287 | \\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 288 | \<body>\n\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 289 | \<div class=\"head\">\ | 
| 
3f697c636fa1
eliminated fragile HTML.with_charset -- always use utf-8;
 wenzelm parents: 
39616diff
changeset | 290 | \<h1>" ^ plain title ^ "</h1>\n"; | 
| 6324 | 291 | |
| 14541 | 292 | val end_document = "\n</div>\n</body>\n</html>\n"; | 
| 6324 | 293 | |
| 6754 | 294 | fun gen_link how (path, name) = para (href_path path how ^ " to index of " ^ plain name); | 
| 295 | val up_link = gen_link "Up"; | |
| 296 | val back_link = gen_link "Back"; | |
| 6338 | 297 | |
| 298 | ||
| 299 | (* session index *) | |
| 300 | ||
| 17080 | 301 | fun begin_index up (index_path, session) docs graph = | 
| 6754 | 302 |   begin_document ("Index of " ^ session) ^ up_link up ^
 | 
| 303 |   para ("View " ^ href_path graph "theory dependencies" ^
 | |
| 27829 | 304 | implode (map (fn (p, name) => "<br/>\nView " ^ href_path p name) docs)) ^ | 
| 37941 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 305 | "\n</div>\n<div class=\"theories\">\n<h2>Theories</h2>\n<ul>\n"; | 
| 6338 | 306 | |
| 6649 | 307 | fun choice chs s = space_implode " " (map (fn (s', lnk) => | 
| 308 | enclose "[" "]" (if s = s' then keyword s' else href_name lnk s')) chs); | |
| 309 | ||
| 9415 | 310 | fun applet_pages session back = | 
| 6649 | 311 | let | 
| 9415 | 312 | val sizes = | 
| 313 |      [("small", "small.html", ("500", "400")),
 | |
| 314 |       ("medium", "medium.html", ("650", "520")),
 | |
| 315 |       ("large", "large.html", ("800", "640"))];
 | |
| 6649 | 316 | |
| 9415 | 317 | fun applet_page (size, name, (width, height)) = | 
| 6754 | 318 | let | 
| 319 | val browser_size = "Set browser size: " ^ | |
| 9415 | 320 | choice (map (fn (y, z, _) => (y, z)) sizes) size; | 
| 6754 | 321 | in | 
| 322 |         (name, begin_document ("Theory dependencies of " ^ session) ^
 | |
| 323 | back_link back ^ | |
| 9415 | 324 | para browser_size ^ | 
| 37941 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 325 | "\n</div>\n<div class=\"graphbrowser\">\n\ | 
| 17080 | 326 | \<applet code=\"GraphBrowser/GraphBrowser.class\" \ | 
| 14541 | 327 | \archive=\"GraphBrowser.jar\" \ | 
| 27832 
992c6d984925
<applet>: more XHTML 1.0 Transitional conformance;
 wenzelm parents: 
27830diff
changeset | 328 | \width=" ^ quote width ^ " height=" ^ quote height ^ ">\n\ | 
| 
992c6d984925
<applet>: more XHTML 1.0 Transitional conformance;
 wenzelm parents: 
27830diff
changeset | 329 | \<param name=\"graphfile\" value=\"" ^ "session.graph" ^ "\"/>\n\ | 
| 37941 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 330 | \</applet>" ^ end_document) | 
| 6649 | 331 | end; | 
| 9415 | 332 | in map applet_page sizes end; | 
| 6649 | 333 | |
| 334 | ||
| 27829 | 335 | fun entry (p, s) = "<li>" ^ href_path p (plain s) ^ "</li>\n"; | 
| 6405 | 336 | |
| 6338 | 337 | val theory_entry = entry; | 
| 338 | ||
| 6405 | 339 | fun session_entries [] = "</ul>" | 
| 340 | | session_entries es = | |
| 37941 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 341 | "</ul>\n</div>\n<div class=\"sessions\">\n\ | 
| 17080 | 342 | \<h2>Sessions</h2>\n<ul>\n" ^ implode (map entry es) ^ "</ul>"; | 
| 6324 | 343 | |
| 344 | ||
| 345 | (* theory *) | |
| 346 | ||
| 27861 
911bf8e58c4c
removed obsolete verbatim_source, results, chapter, section etc.;
 wenzelm parents: 
27832diff
changeset | 347 | val theory_source = enclose | 
| 37941 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 348 | "\n</div>\n<div class=\"source\">\n<pre>" | 
| 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 349 | "</pre>\n"; | 
| 6324 | 350 | |
| 351 | ||
| 352 | local | |
| 27490 | 353 | fun file (href, path, loadit) = | 
| 354 |     href_path href (if loadit then file_path path else enclose "(" ")" (file_path path));
 | |
| 6361 | 355 | |
| 6376 | 356 | fun theory up A = | 
| 357 |     begin_document ("Theory " ^ A) ^ "\n" ^ up_link up ^
 | |
| 23622 | 358 | command "theory" ^ " " ^ name A; | 
| 16267 
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
 wenzelm parents: 
16196diff
changeset | 359 | |
| 
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
 wenzelm parents: 
16196diff
changeset | 360 | fun imports Bs = | 
| 
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
 wenzelm parents: 
16196diff
changeset | 361 | keyword "imports" ^ " " ^ space_implode " " (map (uncurry href_opt_path o apsnd name) Bs); | 
| 
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
 wenzelm parents: 
16196diff
changeset | 362 | |
| 27829 | 363 | fun uses Ps = keyword "uses" ^ " " ^ space_implode " " (map file Ps) ^ "<br/>\n"; | 
| 6324 | 364 | in | 
| 365 | ||
| 16267 
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
 wenzelm parents: 
16196diff
changeset | 366 | fun begin_theory up A Bs Ps body = | 
| 27829 | 367 | theory up A ^ "<br/>\n" ^ | 
| 368 | imports Bs ^ "<br/>\n" ^ | |
| 16267 
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
 wenzelm parents: 
16196diff
changeset | 369 | (if null Ps then "" else uses Ps) ^ | 
| 
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
 wenzelm parents: 
16196diff
changeset | 370 | body; | 
| 
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
 wenzelm parents: 
16196diff
changeset | 371 | |
| 6324 | 372 | end; | 
| 373 | ||
| 374 | ||
| 37940 
4857eab31298
generic external source files -- nothing special about ML here;
 wenzelm parents: 
37146diff
changeset | 375 | (* external file *) | 
| 6324 | 376 | |
| 37940 
4857eab31298
generic external source files -- nothing special about ML here;
 wenzelm parents: 
37146diff
changeset | 377 | fun external_file path str = | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
20742diff
changeset | 378 |   begin_document ("File " ^ Url.implode path) ^
 | 
| 37941 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 379 | "\n</div><div class=\"external_source\">\n" ^ | 
| 14571 | 380 | verbatim str ^ | 
| 37941 
1d812ff95a14
refrain from generating <hr/> and from "hiding" it in isabelle.css -- the latter might be used in other situations as well;
 wenzelm parents: 
37940diff
changeset | 381 | "\n</div>\n<div class=\"external_footer\">" ^ | 
| 14541 | 382 | end_document; | 
| 6324 | 383 | |
| 7572 | 384 | end; |