author | kleing |
Mon, 21 Jun 2004 10:25:57 +0200 | |
changeset 14981 | e73f8140af78 |
parent 14899 | d9b6c81b52ac |
child 15349 | 440687010501 |
permissions | -rw-r--r-- |
9140 | 1 |
(* Title: Pure/Isar/isar_output.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
5 |
Isar theory output. |
|
6 |
*) |
|
7 |
||
8 |
signature ISAR_OUTPUT = |
|
9 |
sig |
|
14899 | 10 |
val display: bool ref |
11 |
val quotes: bool ref |
|
12 |
val indent: int ref |
|
13 |
val source: bool ref |
|
10321 | 14 |
val add_commands: (string * (Args.src -> Toplevel.state -> string)) list -> unit |
9140 | 15 |
val add_options: (string * (string -> (unit -> string) -> unit -> string)) list -> unit |
9220 | 16 |
val print_antiquotations: unit -> unit |
9140 | 17 |
val boolean: string -> bool |
18 |
val integer: string -> int |
|
19 |
val args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> |
|
10321 | 20 |
(Args.src -> Proof.context -> 'a -> string) -> Args.src -> Toplevel.state -> string |
9140 | 21 |
datatype markup = Markup | MarkupEnv | Verbatim |
22 |
val interest_level: int ref |
|
11714 | 23 |
val modes: string list ref |
12939 | 24 |
val eval_antiquote: Scan.lexicon -> Toplevel.state -> string * Position.T -> string |
9140 | 25 |
val parse_thy: (markup -> OuterLex.token list -> OuterLex.token * OuterLex.token list) -> |
26 |
Scan.lexicon -> Toplevel.transition list -> (OuterLex.token, 'a) Source.source -> |
|
27 |
(Toplevel.transition * (Toplevel.state -> Buffer.T -> Buffer.T)) list * Buffer.T |
|
11716 | 28 |
val output_with: (Proof.context -> 'a -> Pretty.T) -> Args.src -> |
29 |
Proof.context -> 'a -> string |
|
9140 | 30 |
end; |
31 |
||
32 |
structure IsarOutput: ISAR_OUTPUT = |
|
33 |
struct |
|
34 |
||
35 |
structure T = OuterLex; |
|
36 |
structure P = OuterParse; |
|
37 |
||
38 |
||
14899 | 39 |
|
40 |
(** global references -- defaults for options **) |
|
41 |
||
42 |
val locale = ref ""; |
|
43 |
val display = ref false; |
|
44 |
val quotes = ref false; |
|
45 |
val indent = ref 0; |
|
46 |
val source = ref false; |
|
47 |
val break = ref false; |
|
48 |
||
49 |
||
50 |
||
9140 | 51 |
(** maintain global commands **) |
52 |
||
53 |
local |
|
54 |
||
55 |
val global_commands = |
|
10321 | 56 |
ref (Symtab.empty: (Args.src -> Toplevel.state -> string) Symtab.table); |
9140 | 57 |
|
58 |
val global_options = |
|
59 |
ref (Symtab.empty: (string -> (unit -> string) -> unit -> string) Symtab.table); |
|
60 |
||
61 |
||
62 |
fun add_item kind (tab, (name, x)) = |
|
63 |
(if is_none (Symtab.lookup (tab, name)) then () |
|
64 |
else warning ("Redefined antiquotation output " ^ kind ^ ": " ^ quote name); |
|
65 |
Symtab.update ((name, x), tab)); |
|
66 |
||
67 |
fun add_items kind xs tab = foldl (add_item kind) (tab, xs); |
|
68 |
||
69 |
in |
|
70 |
||
71 |
val add_commands = Library.change global_commands o (add_items "command"); |
|
72 |
val add_options = Library.change global_options o (add_items "option"); |
|
73 |
||
74 |
fun command src = |
|
75 |
let val ((name, _), pos) = Args.dest_src src in |
|
76 |
(case Symtab.lookup (! global_commands, name) of |
|
77 |
None => error ("Unknown antiquotation command: " ^ quote name ^ Position.str_of pos) |
|
12881 | 78 |
| Some f => transform_failure (curry Antiquote.ANTIQUOTE_FAIL (name, pos)) (f src)) |
9140 | 79 |
end; |
80 |
||
81 |
fun option (name, s) f () = |
|
82 |
(case Symtab.lookup (! global_options, name) of |
|
83 |
None => error ("Unknown antiquotation option: " ^ quote name) |
|
84 |
| Some opt => opt s f ()); |
|
85 |
||
86 |
fun options [] f = f |
|
87 |
| options (opt :: opts) f = option opt (options opts f); |
|
88 |
||
9213 | 89 |
|
9220 | 90 |
fun print_antiquotations () = |
91 |
[Pretty.big_list "antiquotation commands:" (map Pretty.str (Symtab.keys (! global_commands))), |
|
92 |
Pretty.big_list "antiquotation options:" (map Pretty.str (Symtab.keys (! global_options)))] |
|
93 |
|> Pretty.chunks |> Pretty.writeln; |
|
9213 | 94 |
|
9140 | 95 |
end; |
96 |
||
97 |
||
98 |
||
99 |
(** syntax of antiquotations **) |
|
100 |
||
101 |
(* option values *) |
|
102 |
||
103 |
fun boolean "" = true |
|
104 |
| boolean "true" = true |
|
105 |
| boolean "false" = false |
|
106 |
| boolean s = error ("Bad boolean value: " ^ quote s); |
|
107 |
||
108 |
fun integer s = |
|
109 |
let |
|
110 |
fun int ss = |
|
14899 | 111 |
(case Library.read_int ss of (i, []) => i |
112 |
| _ => error ("Bad integer value: " ^ quote s)); |
|
9140 | 113 |
in (case Symbol.explode s of "-" :: ss => ~ (int ss) | ss => int ss) end; |
114 |
||
115 |
||
116 |
(* args syntax *) |
|
117 |
||
118 |
fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) = |
|
119 |
Args.syntax "antiquotation" scan; |
|
120 |
||
10321 | 121 |
fun args scan f src state : string = |
14899 | 122 |
let |
123 |
val ctxt0 = |
|
124 |
if ! locale = "" then Toplevel.context_of state |
|
125 |
else #3 (Locale.read_context_statement (Some (! locale)) [] [] |
|
126 |
(ProofContext.init (Toplevel.theory_of state))); |
|
127 |
val (ctxt, x) = syntax scan src ctxt0; |
|
10321 | 128 |
in f src ctxt x end; |
9140 | 129 |
|
130 |
||
131 |
(* outer syntax *) |
|
132 |
||
133 |
local |
|
134 |
||
135 |
val property = P.xname -- Scan.optional (P.$$$ "=" |-- P.!!! P.xname) ""; |
|
136 |
val properties = Scan.optional (P.$$$ "[" |-- P.!!! (P.enum "," property --| P.$$$ "]")) []; |
|
137 |
||
138 |
val antiq = P.position P.xname -- properties -- P.arguments --| Scan.ahead P.eof |
|
139 |
>> (fn (((x, pos), y), z) => (y, Args.src ((x, z), pos))); |
|
140 |
||
141 |
fun antiq_args_aux keyword_lexicon (str, pos) = |
|
142 |
Source.of_string str |
|
143 |
|> Symbol.source false |
|
144 |
|> T.source false (K (keyword_lexicon, Scan.empty_lexicon)) pos |
|
145 |
|> T.source_proper |
|
146 |
|> Source.source T.stopper (Scan.error (Scan.bulk (P.!!! antiq))) None |
|
147 |
|> Source.exhaust; |
|
148 |
||
149 |
in |
|
150 |
||
151 |
fun antiq_args lex (s, pos) = |
|
152 |
(case antiq_args_aux lex (s, pos) of [x] => x | _ => raise ERROR) handle ERROR => |
|
153 |
error ("Malformed antiquotation: " ^ quote ("@{" ^ s ^ "}") ^ Position.str_of pos); |
|
154 |
||
155 |
end; |
|
156 |
||
157 |
||
158 |
(* eval_antiquote *) |
|
159 |
||
11714 | 160 |
val modes = ref ([]: string list); |
161 |
||
9140 | 162 |
fun eval_antiquote lex state (str, pos) = |
163 |
let |
|
164 |
fun expand (Antiquote.Text s) = s |
|
165 |
| expand (Antiquote.Antiq x) = |
|
166 |
let val (opts, src) = antiq_args lex x in |
|
10739 | 167 |
options opts (fn () => command src state) (); (*preview errors!*) |
11714 | 168 |
Library.setmp print_mode (! modes @ Latex.modes @ ! print_mode) |
10570 | 169 |
(options opts (fn () => command src state)) () |
10321 | 170 |
end; |
9140 | 171 |
val ants = Antiquote.antiquotes_of (str, pos); |
172 |
in |
|
173 |
if Toplevel.is_toplevel state andalso exists Antiquote.is_antiq ants then |
|
174 |
error ("Cannot expand antiquotations at top-level" ^ Position.str_of pos) |
|
175 |
else implode (map expand ants) |
|
176 |
end; |
|
177 |
||
178 |
||
179 |
||
180 |
(** present theory source **) |
|
181 |
||
182 |
(* present_tokens *) |
|
183 |
||
184 |
val interest_level = ref 0; |
|
185 |
||
11861 | 186 |
fun present_tokens lex (flag, toks) state = |
187 |
Buffer.add (case flag of None => "" | Some b => Latex.flag_markup b) o |
|
188 |
(toks |
|
189 |
|> mapfilter (fn (tk, i) => if i > ! interest_level then None else Some tk) |
|
190 |
|> map (apsnd (eval_antiquote lex state)) |
|
191 |
|> Latex.output_tokens |
|
192 |
|> Buffer.add); |
|
9140 | 193 |
|
194 |
||
195 |
(* parse_thy *) |
|
196 |
||
197 |
datatype markup = Markup | MarkupEnv | Verbatim; |
|
198 |
||
199 |
local |
|
200 |
||
201 |
val opt_newline = Scan.option (Scan.one T.is_newline); |
|
202 |
||
203 |
fun check_level i = |
|
204 |
if i > 0 then Scan.succeed () |
|
205 |
else Scan.fail_with (K "Bad nesting of meta-comments"); |
|
206 |
||
207 |
val ignore = |
|
13775
a918c547cd4d
corrected swallowing of newlines after end-of-ignore: rollback
oheimb
parents:
13774
diff
changeset
|
208 |
Scan.depend (fn i => opt_newline |-- P.position (Scan.one T.is_begin_ignore) >> pair (i + 1)) || |
a918c547cd4d
corrected swallowing of newlines after end-of-ignore: rollback
oheimb
parents:
13774
diff
changeset
|
209 |
Scan.depend (fn i => P.position (Scan.one T.is_end_ignore) --| |
13774
77a1e723151d
corrected swallowing of newlines after end-of-ignore (improved)
oheimb
parents:
13742
diff
changeset
|
210 |
(opt_newline -- check_level i) >> pair (i - 1)); |
9140 | 211 |
|
212 |
val ignore_cmd = Scan.state -- ignore |
|
11861 | 213 |
>> (fn (i, (x, pos)) => (false, (None, ((Latex.Basic x, ("", pos)), i)))); |
9140 | 214 |
|
215 |
||
216 |
val is_improper = not o (T.is_proper orf T.is_begin_ignore orf T.is_end_ignore); |
|
217 |
val improper = Scan.any is_improper; |
|
218 |
||
219 |
val improper_keep_indent = Scan.repeat |
|
220 |
(Scan.unless (Scan.one T.is_indent -- Scan.one T.is_proper) (Scan.one is_improper)); |
|
221 |
||
222 |
val improper_end = |
|
223 |
(improper -- P.semicolon) |-- improper_keep_indent || |
|
224 |
improper_keep_indent; |
|
225 |
||
226 |
||
227 |
val stopper = |
|
11861 | 228 |
((false, (None, ((Latex.Basic (#1 T.stopper), ("", Position.none)), 0))), |
229 |
fn (_, (_, ((Latex.Basic x, _), _))) => (#2 T.stopper x) | _ => false); |
|
9140 | 230 |
|
231 |
in |
|
232 |
||
233 |
fun parse_thy markup lex trs src = |
|
234 |
let |
|
235 |
val text = P.position P.text; |
|
9220 | 236 |
val token = Scan.depend (fn i => |
9140 | 237 |
(improper |-- markup Markup -- P.!!!! (improper |-- text --| improper_end) |
11861 | 238 |
>> (fn (x, y) => (true, (Some true, ((Latex.Markup (T.val_of x), y), i)))) || |
9140 | 239 |
improper |-- markup MarkupEnv -- P.!!!! (improper |-- text --| improper_end) |
11861 | 240 |
>> (fn (x, y) => (true, (Some true, ((Latex.MarkupEnv (T.val_of x), y), i)))) || |
9140 | 241 |
(P.$$$ "--" |-- P.!!!! (improper |-- text)) |
11861 | 242 |
>> (fn y => (false, (None, ((Latex.Markup "cmt", y), i)))) || |
9140 | 243 |
(improper -- markup Verbatim) |-- P.!!!! (improper |-- text --| improper_end) |
11861 | 244 |
>> (fn y => (true, (None, ((Latex.Verbatim, y), i)))) || |
9140 | 245 |
P.position (Scan.one T.not_eof) |
11861 | 246 |
>> (fn (x, pos) => (T.is_kind T.Command x, (Some false, ((Latex.Basic x, ("", pos)), i))))) |
9220 | 247 |
>> pair i); |
9140 | 248 |
|
249 |
val body = Scan.any (not o fst andf not o #2 stopper); |
|
11861 | 250 |
val section = body -- Scan.one fst -- body |
251 |
>> (fn ((b1, c), b2) => (#1 (#2 c), map (snd o snd) (b1 @ (c :: b2)))); |
|
9140 | 252 |
|
253 |
val cmds = |
|
254 |
src |
|
255 |
|> Source.filter (not o T.is_semicolon) |
|
256 |
|> Source.source' 0 T.stopper (Scan.error (Scan.bulk (ignore_cmd || token))) None |
|
257 |
|> Source.source stopper (Scan.error (Scan.bulk section)) None |
|
258 |
|> Source.exhaust; |
|
259 |
in |
|
260 |
if length cmds = length trs then |
|
261 |
(trs ~~ map (present_tokens lex) cmds, Buffer.empty) |
|
11861 | 262 |
else error "Messed-up outer syntax for presentation" |
9140 | 263 |
end; |
264 |
||
265 |
end; |
|
266 |
||
267 |
||
268 |
||
269 |
(** setup default output **) |
|
270 |
||
271 |
(* options *) |
|
272 |
||
273 |
val _ = add_options |
|
9220 | 274 |
[("show_types", Library.setmp Syntax.show_types o boolean), |
275 |
("show_sorts", Library.setmp Syntax.show_sorts o boolean), |
|
14707 | 276 |
("show_structs", Library.setmp show_structs o boolean), |
14696 | 277 |
("long_names", Library.setmp NameSpace.long_names o boolean), |
9220 | 278 |
("eta_contract", Library.setmp Syntax.eta_contract o boolean), |
14899 | 279 |
("locale", Library.setmp locale), |
9140 | 280 |
("display", Library.setmp display o boolean), |
13929
21615e44ba88
Added "break" flag to allow line breaks within \isa{...}
berghofe
parents:
13775
diff
changeset
|
281 |
("break", Library.setmp break o boolean), |
9140 | 282 |
("quotes", Library.setmp quotes o boolean), |
283 |
("mode", fn s => fn f => fn () => Library.setmp print_mode (s :: ! print_mode) f ()), |
|
9732 | 284 |
("margin", Pretty.setmp_margin o integer), |
9750 | 285 |
("indent", Library.setmp indent o integer), |
10321 | 286 |
("source", Library.setmp source o boolean), |
287 |
("goals_limit", Library.setmp goals_limit o integer)]; |
|
9140 | 288 |
|
289 |
||
290 |
(* commands *) |
|
291 |
||
292 |
fun cond_quote prt = |
|
293 |
if ! quotes then Pretty.quote prt else prt; |
|
294 |
||
295 |
fun cond_display prt = |
|
296 |
if ! display then |
|
14835 | 297 |
Output.output (Pretty.string_of (Pretty.indent (! indent) prt)) |
9863 | 298 |
|> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}" |
9140 | 299 |
else |
14835 | 300 |
Output.output ((if ! break then Pretty.string_of else Pretty.str_of) prt) |
9140 | 301 |
|> enclose "\\isa{" "}"; |
302 |
||
14345 | 303 |
fun cond_seq_display prts = |
304 |
if ! display then |
|
14835 | 305 |
map (Output.output o Pretty.string_of o Pretty.indent (! indent)) prts |
14345 | 306 |
|> separate "\\isasep\\isanewline%\n" |
307 |
|> implode |
|
308 |
|> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}" |
|
309 |
else |
|
14835 | 310 |
map (Output.output o (if ! break then Pretty.string_of else Pretty.str_of)) prts |
14345 | 311 |
|> separate "\\isasep\\isanewline%\n" |
312 |
|> implode |
|
313 |
|> enclose "\\isa{" "}"; |
|
314 |
||
11011
14b78c0c9f05
pretty_text: tweak_lines handles linebreaks gracefully;
wenzelm
parents:
10739
diff
changeset
|
315 |
fun tweak_line s = |
14b78c0c9f05
pretty_text: tweak_lines handles linebreaks gracefully;
wenzelm
parents:
10739
diff
changeset
|
316 |
if ! display then s else Symbol.strip_blanks s; |
9750 | 317 |
|
11011
14b78c0c9f05
pretty_text: tweak_lines handles linebreaks gracefully;
wenzelm
parents:
10739
diff
changeset
|
318 |
val pretty_text = Pretty.chunks o map Pretty.str o map tweak_line o Library.split_lines; |
9140 | 319 |
|
9750 | 320 |
val pretty_source = |
321 |
pretty_text o space_implode " " o map Args.string_of o #2 o #1 o Args.dest_src; |
|
9140 | 322 |
|
12055 | 323 |
fun pretty_term ctxt = ProofContext.pretty_term ctxt o ProofContext.extern_skolem ctxt; |
9140 | 324 |
|
14345 | 325 |
fun pretty_thm ctxt = pretty_term ctxt o Thm.prop_of; |
9140 | 326 |
|
12055 | 327 |
fun pretty_prf full ctxt thms = (* FIXME context syntax!? *) |
11524
197f2e14a714
Added functions for printing primitive proof terms.
berghofe
parents:
11240
diff
changeset
|
328 |
Pretty.chunks (map (ProofSyntax.pretty_proof_of full) thms); |
197f2e14a714
Added functions for printing primitive proof terms.
berghofe
parents:
11240
diff
changeset
|
329 |
|
9750 | 330 |
fun output_with pretty src ctxt x = |
331 |
let |
|
12055 | 332 |
val prt = pretty ctxt x; (*always pretty in order to catch errors!*) |
9750 | 333 |
val prt' = if ! source then pretty_source src else prt; |
334 |
in cond_display (cond_quote prt') end; |
|
9732 | 335 |
|
14345 | 336 |
fun output_seq_with pretty src ctxt xs = |
337 |
let |
|
338 |
val prts = map (pretty ctxt) xs; (*always pretty in order to catch errors!*) |
|
339 |
val prts' = if ! source then [pretty_source src] else prts; |
|
340 |
in cond_seq_display (map cond_quote prts') end; |
|
341 |
||
10360 | 342 |
fun output_goals main_goal src state = args (Scan.succeed ()) (output_with (fn _ => fn _ => |
343 |
Pretty.chunks (Proof.pretty_goals main_goal (Toplevel.proof_of state |
|
344 |
handle Toplevel.UNDEF => error "No proof state")))) src state; |
|
345 |
||
9140 | 346 |
val _ = add_commands |
14696 | 347 |
[("thm", args Attrib.local_thmss (output_seq_with pretty_thm)), |
9750 | 348 |
("prop", args Args.local_prop (output_with pretty_term)), |
349 |
("term", args Args.local_term (output_with pretty_term)), |
|
14776 | 350 |
("typ", args Args.local_typ_raw (output_with ProofContext.pretty_typ)), |
14696 | 351 |
("text", args (Scan.lift Args.name) (output_with (K pretty_text))), |
10360 | 352 |
("goals", output_goals true), |
14696 | 353 |
("subgoals", output_goals false), |
354 |
("prf", args Attrib.local_thmss (output_with (pretty_prf false))), |
|
355 |
("full_prf", args Attrib.local_thmss (output_with (pretty_prf true)))]; |
|
9140 | 356 |
|
357 |
end; |