author | wenzelm |
Fri, 19 Jan 2007 22:08:31 +0100 | |
changeset 22122 | 58f846cc5c3d |
parent 21858 | 05f57309170c |
child 22796 | 34c316d7b630 |
permissions | -rw-r--r-- |
6324 | 1 |
(* Title: Pure/Thy/HTML.ML |
2 |
ID: $Id$ |
|
9415 | 3 |
Author: Markus Wenzel and Stefan Berghofer, TU Muenchen |
6324 | 4 |
|
9415 | 5 |
HTML presentation elements. |
6324 | 6 |
*) |
7 |
||
8 |
signature HTML = |
|
9 |
sig |
|
10 |
val output: string -> string |
|
11 |
val output_width: string -> string * real |
|
12 |
type text (* = string *) |
|
13 |
val plain: string -> text |
|
14 |
val name: string -> text |
|
15 |
val keyword: string -> text |
|
16 |
val file_name: string -> text |
|
6649 | 17 |
val file_path: Url.T -> text |
6324 | 18 |
val href_name: string -> text -> text |
6649 | 19 |
val href_path: Url.T -> text -> text |
6376 | 20 |
val href_opt_name: string option -> text -> text |
6649 | 21 |
val href_opt_path: Url.T option -> text -> text |
6376 | 22 |
val para: text -> text |
6324 | 23 |
val preform: text -> text |
24 |
val verbatim: string -> text |
|
17470
6e9d910c3837
added with_charset: string -> ('a -> 'b) -> 'a -> 'b;
wenzelm
parents:
17412
diff
changeset
|
25 |
val with_charset: string -> ('a -> 'b) -> 'a -> 'b |
17080 | 26 |
val begin_index: Url.T * string -> Url.T * string -> (Url.T * string) list -> Url.T -> text |
6338 | 27 |
val end_index: text |
9415 | 28 |
val applet_pages: string -> Url.T * string -> (string * string) list |
6649 | 29 |
val theory_entry: Url.T * string -> text |
30 |
val session_entries: (Url.T * string) list -> text |
|
8190 | 31 |
val verbatim_source: Symbol.symbol list -> text |
6649 | 32 |
val begin_theory: Url.T * string -> string -> (Url.T option * string) list -> |
7408 | 33 |
(Url.T option * Url.T * bool option) list -> text -> text |
6324 | 34 |
val end_theory: text |
6649 | 35 |
val ml_file: Url.T -> string -> text |
22122 | 36 |
val results: Proof.context -> string -> (string * thm list) list -> text |
8088 | 37 |
val chapter: string -> text |
6324 | 38 |
val section: string -> text |
7634 | 39 |
val subsection: string -> text |
40 |
val subsubsection: string -> text |
|
6324 | 41 |
end; |
42 |
||
43 |
structure HTML: HTML = |
|
44 |
struct |
|
45 |
||
46 |
||
47 |
(** HTML print modes **) |
|
48 |
||
49 |
(* mode *) |
|
50 |
||
51 |
val htmlN = "HTML"; |
|
10573 | 52 |
fun html_mode f x = setmp print_mode (htmlN :: ! print_mode) f x; |
6324 | 53 |
|
54 |
||
55 |
(* symbol output *) |
|
56 |
||
6338 | 57 |
local |
16196 | 58 |
val html_syms = Symtab.make |
59 |
[("\\<spacespace>", (2.0, " ")), |
|
60 |
("\\<exclamdown>", (1.0, "¡")), |
|
61 |
("\\<cent>", (1.0, "¢")), |
|
62 |
("\\<pounds>", (1.0, "£")), |
|
63 |
("\\<currency>", (1.0, "¤")), |
|
64 |
("\\<yen>", (1.0, "¥")), |
|
65 |
("\\<bar>", (1.0, "¦")), |
|
66 |
("\\<section>", (1.0, "§")), |
|
67 |
("\\<dieresis>", (1.0, "¨")), |
|
68 |
("\\<copyright>", (1.0, "©")), |
|
69 |
("\\<ordfeminine>", (1.0, "ª")), |
|
70 |
("\\<guillemotleft>", (1.0, "«")), |
|
71 |
("\\<not>", (1.0, "¬")), |
|
72 |
("\\<hyphen>", (1.0, "­")), |
|
73 |
("\\<registered>", (1.0, "®")), |
|
74 |
("\\<inverse>", (1.0, "¯")), |
|
75 |
("\\<degree>", (1.0, "°")), |
|
76 |
("\\<plusminus>", (1.0, "±")), |
|
77 |
("\\<twosuperior>", (1.0, "²")), |
|
78 |
("\\<threesuperior>", (1.0, "³")), |
|
79 |
("\\<acute>", (1.0, "´")), |
|
80 |
("\\<paragraph>", (1.0, "¶")), |
|
81 |
("\\<cdot>", (1.0, "·")), |
|
82 |
("\\<cedilla>", (1.0, "¸")), |
|
83 |
("\\<onesuperior>", (1.0, "¹")), |
|
84 |
("\\<ordmasculine>", (1.0, "º")), |
|
85 |
("\\<guillemotright>", (1.0, "»")), |
|
86 |
("\\<onequarter>", (1.0, "¼")), |
|
87 |
("\\<onehalf>", (1.0, "½")), |
|
88 |
("\\<threequarters>", (1.0, "¾")), |
|
89 |
("\\<questiondown>", (1.0, "¿")), |
|
90 |
("\\<times>", (1.0, "×")), |
|
91 |
("\\<div>", (1.0, "÷")), |
|
92 |
("\\<circ>", (1.0, "o")), |
|
93 |
("\\<Alpha>", (1.0, "Α")), |
|
94 |
("\\<Beta>", (1.0, "Β")), |
|
95 |
("\\<Gamma>", (1.0, "Γ")), |
|
96 |
("\\<Delta>", (1.0, "Δ")), |
|
97 |
("\\<Epsilon>", (1.0, "Ε")), |
|
98 |
("\\<Zeta>", (1.0, "Ζ")), |
|
99 |
("\\<Eta>", (1.0, "Η")), |
|
100 |
("\\<Theta>", (1.0, "Θ")), |
|
101 |
("\\<Iota>", (1.0, "Ι")), |
|
102 |
("\\<Kappa>", (1.0, "Κ")), |
|
103 |
("\\<Lambda>", (1.0, "Λ")), |
|
104 |
("\\<Mu>", (1.0, "Μ")), |
|
105 |
("\\<Nu>", (1.0, "Ν")), |
|
106 |
("\\<Xi>", (1.0, "Ξ")), |
|
107 |
("\\<Omicron>", (1.0, "Ο")), |
|
108 |
("\\<Pi>", (1.0, "Π")), |
|
109 |
("\\<Rho>", (1.0, "Ρ")), |
|
110 |
("\\<Sigma>", (1.0, "Σ")), |
|
111 |
("\\<Tau>", (1.0, "Τ")), |
|
112 |
("\\<Upsilon>", (1.0, "Υ")), |
|
113 |
("\\<Phi>", (1.0, "Φ")), |
|
114 |
("\\<Chi>", (1.0, "Χ")), |
|
115 |
("\\<Psi>", (1.0, "Ψ")), |
|
116 |
("\\<Omega>", (1.0, "Ω")), |
|
117 |
("\\<alpha>", (1.0, "α")), |
|
118 |
("\\<beta>", (1.0, "β")), |
|
119 |
("\\<gamma>", (1.0, "γ")), |
|
120 |
("\\<delta>", (1.0, "δ")), |
|
121 |
("\\<epsilon>", (1.0, "ε")), |
|
122 |
("\\<zeta>", (1.0, "ζ")), |
|
123 |
("\\<eta>", (1.0, "η")), |
|
124 |
("\\<theta>", (1.0, "ϑ")), |
|
125 |
("\\<iota>", (1.0, "ι")), |
|
126 |
("\\<kappa>", (1.0, "κ")), |
|
127 |
("\\<lambda>", (1.0, "λ")), |
|
128 |
("\\<mu>", (1.0, "μ")), |
|
129 |
("\\<nu>", (1.0, "ν")), |
|
130 |
("\\<xi>", (1.0, "ξ")), |
|
131 |
("\\<omicron>", (1.0, "ο")), |
|
132 |
("\\<pi>", (1.0, "π")), |
|
133 |
("\\<rho>", (1.0, "ρ")), |
|
134 |
("\\<sigma>", (1.0, "σ")), |
|
135 |
("\\<tau>", (1.0, "τ")), |
|
136 |
("\\<upsilon>", (1.0, "υ")), |
|
137 |
("\\<phi>", (1.0, "φ")), |
|
138 |
("\\<chi>", (1.0, "χ")), |
|
139 |
("\\<psi>", (1.0, "ψ")), |
|
140 |
("\\<omega>", (1.0, "ω")), |
|
141 |
("\\<bullet>", (1.0, "•")), |
|
142 |
("\\<dots>", (1.0, "…")), |
|
143 |
("\\<Re>", (1.0, "ℜ")), |
|
144 |
("\\<Im>", (1.0, "ℑ")), |
|
145 |
("\\<wp>", (1.0, "℘")), |
|
146 |
("\\<forall>", (1.0, "∀")), |
|
147 |
("\\<partial>", (1.0, "∂")), |
|
148 |
("\\<exists>", (1.0, "∃")), |
|
149 |
("\\<emptyset>", (1.0, "∅")), |
|
150 |
("\\<nabla>", (1.0, "∇")), |
|
151 |
("\\<in>", (1.0, "∈")), |
|
152 |
("\\<notin>", (1.0, "∉")), |
|
153 |
("\\<Prod>", (1.0, "∏")), |
|
154 |
("\\<Sum>", (1.0, "∑")), |
|
155 |
("\\<star>", (1.0, "∗")), |
|
156 |
("\\<propto>", (1.0, "∝")), |
|
157 |
("\\<infinity>", (1.0, "∞")), |
|
158 |
("\\<angle>", (1.0, "∠")), |
|
159 |
("\\<and>", (1.0, "∧")), |
|
160 |
("\\<or>", (1.0, "∨")), |
|
161 |
("\\<inter>", (1.0, "∩")), |
|
162 |
("\\<union>", (1.0, "∪")), |
|
163 |
("\\<sim>", (1.0, "∼")), |
|
164 |
("\\<cong>", (1.0, "≅")), |
|
165 |
("\\<approx>", (1.0, "≈")), |
|
166 |
("\\<noteq>", (1.0, "≠")), |
|
167 |
("\\<equiv>", (1.0, "≡")), |
|
168 |
("\\<le>", (1.0, "≤")), |
|
169 |
("\\<ge>", (1.0, "≥")), |
|
170 |
("\\<subset>", (1.0, "⊂")), |
|
171 |
("\\<supset>", (1.0, "⊃")), |
|
172 |
("\\<subseteq>", (1.0, "⊆")), |
|
173 |
("\\<supseteq>", (1.0, "⊇")), |
|
174 |
("\\<oplus>", (1.0, "⊕")), |
|
175 |
("\\<otimes>", (1.0, "⊗")), |
|
176 |
("\\<bottom>", (1.0, "⊥")), |
|
177 |
("\\<lceil>", (1.0, "⌈")), |
|
178 |
("\\<rceil>", (1.0, "⌉")), |
|
179 |
("\\<lfloor>", (1.0, "⌊")), |
|
180 |
("\\<rfloor>", (1.0, "⌋")), |
|
181 |
("\\<langle>", (1.0, "⟨")), |
|
182 |
("\\<rangle>", (1.0, "⟩")), |
|
183 |
("\\<lozenge>", (1.0, "◊")), |
|
184 |
("\\<spadesuit>", (1.0, "♠")), |
|
185 |
("\\<clubsuit>", (1.0, "♣")), |
|
186 |
("\\<heartsuit>", (1.0, "♥")), |
|
187 |
("\\<diamondsuit>", (1.0, "♦")), |
|
188 |
("\\<lbrakk>", (2.0, "[|")), |
|
189 |
("\\<rbrakk>", (2.0, "|]")), |
|
190 |
("\\<Longrightarrow>", (3.0, "==>")), |
|
191 |
("\\<Rightarrow>", (2.0, "=>")), |
|
192 |
("\\<And>", (2.0, "!!")), |
|
193 |
("\\<Colon>", (2.0, "::")), |
|
194 |
("\\<lparr>", (2.0, "(|")), |
|
195 |
("\\<rparr>", (2.0, "|)),")), |
|
196 |
("\\<longleftrightarrow>", (3.0, "<->")), |
|
197 |
("\\<longrightarrow>", (3.0, "-->")), |
|
198 |
("\\<rightarrow>", (2.0, "->")), |
|
199 |
("\\<^bsub>", (0.0, "<sub>")), |
|
200 |
("\\<^esub>", (0.0, "</sub>")), |
|
201 |
("\\<^bsup>", (0.0, "<sup>")), |
|
17178 | 202 |
("\\<^esup>", (0.0, "</sup>"))]; |
203 |
||
204 |
val escape = fn "<" => "<" | ">" => ">" | "&" => "&" | c => c; |
|
16196 | 205 |
|
206 |
fun output_sym s = |
|
20742 | 207 |
if Symbol.is_raw s then (1.0, Symbol.decode_raw s) |
208 |
else |
|
209 |
(case Symtab.lookup html_syms s of SOME x => x |
|
210 |
| NONE => (real (size s), translate_string escape s)); |
|
6324 | 211 |
|
17178 | 212 |
fun output_sub s = apsnd (enclose "<sub>" "</sub>") (output_sym s); |
213 |
fun output_sup s = apsnd (enclose "<sup>" "</sup>") (output_sym s); |
|
214 |
fun output_loc s = apsnd (enclose "<span class=\"loc\">" "</span>") (output_sym s); |
|
14534 | 215 |
|
17178 | 216 |
fun output_syms ("\\<^sub>" :: s :: ss) = output_sub s :: output_syms ss |
217 |
| output_syms ("\\<^isub>" :: s :: ss) = output_sub s :: output_syms ss |
|
218 |
| output_syms ("\\<^sup>" :: s :: ss) = output_sup s :: output_syms ss |
|
219 |
| output_syms ("\\<^isup>" :: s :: ss) = output_sup s :: output_syms ss |
|
220 |
| output_syms ("\\<^loc>" :: s :: ss) = output_loc s :: output_syms ss |
|
221 |
| output_syms (s :: ss) = output_sym s :: output_syms ss |
|
222 |
| output_syms [] = []; |
|
6338 | 223 |
in |
6324 | 224 |
|
6338 | 225 |
fun output_width str = |
19305 | 226 |
if not (exists_string (fn s => s = "\\" orelse s = "<" orelse s = ">" orelse s = "&") str) |
14992 | 227 |
then Symbol.default_output str |
6338 | 228 |
else |
17178 | 229 |
let val (syms, width) = fold_map (fn (w, s) => fn width => (s, w + width)) |
230 |
(output_syms (Symbol.explode str)) 0.0 |
|
231 |
in (implode syms, width) end; |
|
6338 | 232 |
|
233 |
val output = #1 o output_width; |
|
17178 | 234 |
val output_symbols = map #2 o output_syms; |
6338 | 235 |
|
236 |
end; |
|
237 |
||
6324 | 238 |
|
239 |
(* token translations *) |
|
240 |
||
16196 | 241 |
fun style s = |
242 |
apfst (enclose ("<span class=" ^ Library.quote s ^ ">") "</span>") o output_width; |
|
6324 | 243 |
|
244 |
val html_trans = |
|
14541 | 245 |
[("class", style "tclass"), |
246 |
("tfree", style "tfree"), |
|
247 |
("tvar", style "tvar"), |
|
248 |
("free", style "free"), |
|
249 |
("bound", style "bound"), |
|
250 |
("var", style "var"), |
|
251 |
("xstr", style "xstr")]; |
|
6324 | 252 |
|
18708 | 253 |
val _ = Context.add_setup (Theory.add_mode_tokentrfuns htmlN html_trans); |
15801 | 254 |
|
6324 | 255 |
|
256 |
||
257 |
(** HTML markup **) |
|
258 |
||
259 |
type text = string; |
|
260 |
||
261 |
||
262 |
(* atoms *) |
|
263 |
||
264 |
val plain = output; |
|
19265 | 265 |
val name = enclose "<span class=\"name\">" "</span>" o output; |
266 |
val keyword = enclose "<span class=\"keyword\">" "</span>" o output; |
|
267 |
val keyword_width = apfst (enclose "<span class=\"keyword\">" "</span>") o output_width; |
|
268 |
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:
20742
diff
changeset
|
269 |
val file_path = file_name o Url.implode; |
6324 | 270 |
|
271 |
||
272 |
(* misc *) |
|
273 |
||
14598
7009f59711e3
Replaced quote by Library.quote, since quote now refers to Symbol.quote
berghofe
parents:
14571
diff
changeset
|
274 |
fun href_name s txt = "<a href=" ^ Library.quote s ^ ">" ^ txt ^ "</a>"; |
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
20742
diff
changeset
|
275 |
fun href_path path txt = href_name (Url.implode path) txt; |
6376 | 276 |
|
15531 | 277 |
fun href_opt_name NONE txt = txt |
278 |
| href_opt_name (SOME s) txt = href_name s txt; |
|
6376 | 279 |
|
15531 | 280 |
fun href_opt_path NONE txt = txt |
281 |
| href_opt_path (SOME p) txt = href_path p txt; |
|
6376 | 282 |
|
12413 | 283 |
fun para txt = "\n<p>" ^ txt ^ "</p>\n"; |
6324 | 284 |
fun preform txt = "<pre>" ^ txt ^ "</pre>"; |
285 |
val verbatim = preform o output; |
|
286 |
||
287 |
||
288 |
(* document *) |
|
289 |
||
17470
6e9d910c3837
added with_charset: string -> ('a -> 'b) -> 'a -> 'b;
wenzelm
parents:
17412
diff
changeset
|
290 |
val charset = ref "iso-8859-1"; |
6e9d910c3837
added with_charset: string -> ('a -> 'b) -> 'a -> 'b;
wenzelm
parents:
17412
diff
changeset
|
291 |
fun with_charset s = setmp charset s; |
6e9d910c3837
added with_charset: string -> ('a -> 'b) -> 'a -> 'b;
wenzelm
parents:
17412
diff
changeset
|
292 |
|
6324 | 293 |
fun begin_document title = |
17080 | 294 |
"<!DOCTYPE HTML PUBLIC \"-//W3C//DTD HTML 4.01 Transitional//EN\" \ |
295 |
\\"http://www.w3.org/TR/html4/loose.dtd\">\n\ |
|
14083
aed5d25c4a0c
Added DOCTYPE and Content-type to HTML documents.
webertj
parents:
13531
diff
changeset
|
296 |
\\n\ |
aed5d25c4a0c
Added DOCTYPE and Content-type to HTML documents.
webertj
parents:
13531
diff
changeset
|
297 |
\<html>\n\ |
6405 | 298 |
\<head>\n\ |
17470
6e9d910c3837
added with_charset: string -> ('a -> 'b) -> 'a -> 'b;
wenzelm
parents:
17412
diff
changeset
|
299 |
\<meta http-equiv=\"Content-Type\" content=\"text/html; charset=" ^ ! charset ^ "\">\n\ |
6405 | 300 |
\<title>" ^ plain (title ^ " (" ^ version ^ ")") ^ "</title>\n\ |
14541 | 301 |
\<link rel=\"stylesheet\" type=\"text/css\" href=\"isabelle.css\">\n\ |
6405 | 302 |
\</head>\n\ |
303 |
\\n\ |
|
304 |
\<body>\n\ |
|
14541 | 305 |
\<div class=\"head\">\ |
6405 | 306 |
\<h1>" ^ plain title ^ "</h1>\n" |
6324 | 307 |
|
14541 | 308 |
val end_document = "\n</div>\n</body>\n</html>\n"; |
6324 | 309 |
|
6754 | 310 |
fun gen_link how (path, name) = para (href_path path how ^ " to index of " ^ plain name); |
311 |
val up_link = gen_link "Up"; |
|
312 |
val back_link = gen_link "Back"; |
|
6338 | 313 |
|
314 |
||
315 |
(* session index *) |
|
316 |
||
17080 | 317 |
fun begin_index up (index_path, session) docs graph = |
6754 | 318 |
begin_document ("Index of " ^ session) ^ up_link up ^ |
319 |
para ("View " ^ href_path graph "theory dependencies" ^ |
|
17080 | 320 |
implode (map (fn (p, name) => "<br>\nView " ^ href_path p name) docs)) ^ |
14541 | 321 |
"\n</div>\n<hr>\n<div class=\"theories\">\n<h2>Theories</h2>\n<ul>\n"; |
6338 | 322 |
|
323 |
val end_index = end_document; |
|
324 |
||
6649 | 325 |
fun choice chs s = space_implode " " (map (fn (s', lnk) => |
326 |
enclose "[" "]" (if s = s' then keyword s' else href_name lnk s')) chs); |
|
327 |
||
9415 | 328 |
fun applet_pages session back = |
6649 | 329 |
let |
9415 | 330 |
val sizes = |
331 |
[("small", "small.html", ("500", "400")), |
|
332 |
("medium", "medium.html", ("650", "520")), |
|
333 |
("large", "large.html", ("800", "640"))]; |
|
6649 | 334 |
|
9415 | 335 |
fun applet_page (size, name, (width, height)) = |
6754 | 336 |
let |
337 |
val browser_size = "Set browser size: " ^ |
|
9415 | 338 |
choice (map (fn (y, z, _) => (y, z)) sizes) size; |
6754 | 339 |
in |
340 |
(name, begin_document ("Theory dependencies of " ^ session) ^ |
|
341 |
back_link back ^ |
|
9415 | 342 |
para browser_size ^ |
17080 | 343 |
"\n</div>\n<hr>\n<div class=\"graphbrowser\">\n\ |
344 |
\<applet code=\"GraphBrowser/GraphBrowser.class\" \ |
|
14541 | 345 |
\archive=\"GraphBrowser.jar\" \ |
9045 | 346 |
\width=" ^ width ^ " height=" ^ height ^ ">\n\ |
9415 | 347 |
\<param name=\"graphfile\" value=\"" ^ "session.graph" ^ "\">\n\ |
6754 | 348 |
\</applet>\n<hr>" ^ end_document) |
6649 | 349 |
end; |
9415 | 350 |
in map applet_page sizes end; |
6649 | 351 |
|
352 |
||
6405 | 353 |
fun entry (p, s) = "<li>" ^ href_path p (plain s) ^ "\n"; |
354 |
||
6338 | 355 |
val theory_entry = entry; |
356 |
||
6405 | 357 |
fun session_entries [] = "</ul>" |
358 |
| session_entries es = |
|
17080 | 359 |
"</ul>\n</div>\n<hr>\n<div class=\"sessions\">\n\ |
360 |
\<h2>Sessions</h2>\n<ul>\n" ^ implode (map entry es) ^ "</ul>"; |
|
6324 | 361 |
|
362 |
||
363 |
(* theory *) |
|
364 |
||
14571 | 365 |
fun verbatim_source ss = |
366 |
"\n</div>\n<hr>\n<div class=\"source\">\n<pre>" ^ |
|
17178 | 367 |
implode (output_symbols ss) ^ |
14571 | 368 |
"</pre>\n</div>\n<hr>\n<div class=\"theorems\">\n"; |
6324 | 369 |
|
370 |
||
371 |
local |
|
15531 | 372 |
fun encl NONE = enclose "[" "]" |
373 |
| encl (SOME false) = enclose "(" ")" |
|
374 |
| encl (SOME true) = I; |
|
6324 | 375 |
|
7408 | 376 |
fun file (opt_ref, path, loadit) = href_opt_path opt_ref (encl loadit (file_path path)); |
6361 | 377 |
|
6376 | 378 |
fun theory up A = |
379 |
begin_document ("Theory " ^ A) ^ "\n" ^ up_link up ^ |
|
16267
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
380 |
keyword "theory" ^ " " ^ name A; |
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
381 |
|
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
382 |
fun imports Bs = |
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
383 |
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:
16196
diff
changeset
|
384 |
|
17209 | 385 |
fun uses Ps = keyword "uses" ^ " " ^ space_implode " " (map file Ps) ^ "<br>\n"; |
6324 | 386 |
in |
387 |
||
16267
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
388 |
fun begin_theory up A Bs Ps body = |
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
389 |
theory up A ^ "<br>\n" ^ |
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
390 |
imports Bs ^ "<br>\n" ^ |
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
391 |
(if null Ps then "" else uses Ps) ^ |
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
392 |
keyword "begin" ^ "<br>\n" ^ |
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
393 |
body; |
0773b17922d8
present new-style theory header, with 'imports' and 'uses';
wenzelm
parents:
16196
diff
changeset
|
394 |
|
6324 | 395 |
end; |
396 |
||
6338 | 397 |
val end_theory = end_document; |
6324 | 398 |
|
399 |
||
400 |
(* ML file *) |
|
401 |
||
402 |
fun ml_file path str = |
|
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
20742
diff
changeset
|
403 |
begin_document ("File " ^ Url.implode path) ^ |
14571 | 404 |
"\n</div>\n<hr><div class=\"mlsource\">\n" ^ |
405 |
verbatim str ^ |
|
406 |
"\n</div>\n<hr>\n<div class=\"mlfooter\">\n" ^ |
|
14541 | 407 |
end_document; |
6324 | 408 |
|
409 |
||
7408 | 410 |
(* theorems *) |
6324 | 411 |
|
7572 | 412 |
local |
413 |
||
22122 | 414 |
fun string_of_thm ctxt = |
14839
c994f1c57fc7
handle raw symbols; Output.add_mode; more robust handling of sub/superscript;
wenzelm
parents:
14775
diff
changeset
|
415 |
Output.output o Pretty.setmp_margin 80 Pretty.string_of o |
22122 | 416 |
setmp show_question_marks false (ProofContext.pretty_thm ctxt); |
6324 | 417 |
|
22122 | 418 |
fun thm ctxt th = preform (prefix_lines " " (html_mode (string_of_thm ctxt) th)); |
7408 | 419 |
|
7572 | 420 |
fun thm_name "" = "" |
421 |
| thm_name a = " " ^ name (a ^ ":"); |
|
422 |
||
423 |
in |
|
424 |
||
22122 | 425 |
fun result ctxt k (a, ths) = para (cat_lines ((keyword k ^ thm_name a) :: map (thm ctxt) ths)); |
7572 | 426 |
|
22122 | 427 |
fun results _ _ [] = "" |
428 |
| results ctxt k (res :: ress) = cat_lines (result ctxt k res :: map (result ctxt "and") ress); |
|
7572 | 429 |
|
430 |
end; |
|
6324 | 431 |
|
432 |
||
7634 | 433 |
(* sections *) |
6324 | 434 |
|
8088 | 435 |
fun chapter heading = "\n\n<h1>" ^ plain heading ^ "</h1>\n"; |
6324 | 436 |
fun section heading = "\n\n<h2>" ^ plain heading ^ "</h2>\n"; |
7634 | 437 |
fun subsection heading = "\n\n<h3>" ^ plain heading ^ "</h3>\n"; |
438 |
fun subsubsection heading = "\n\n<h4>" ^ plain heading ^ "</h4>\n"; |
|
6324 | 439 |
|
440 |
||
19265 | 441 |
(* mode output *) |
442 |
||
443 |
val _ = Output.add_mode htmlN |
|
444 |
(output_width, K keyword_width, Symbol.default_indent, Symbol.encode_raw); |
|
445 |
||
6324 | 446 |
end; |