| author | wenzelm |
| Sat, 05 Jun 2004 13:06:39 +0200 | |
| changeset 14870 | c5cf7c001313 |
| parent 14687 | e089757b952a |
| child 14925 | 0f86a8a694f8 |
| permissions | -rw-r--r-- |
| 5829 | 1 |
(* Title: Pure/Isar/outer_syntax.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
| 8807 | 4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
| 5829 | 5 |
|
6 |
The global Isabelle/Isar outer syntax. |
|
7 |
*) |
|
8 |
||
9 |
signature BASIC_OUTER_SYNTAX = |
|
10 |
sig |
|
11 |
val main: unit -> unit |
|
| 5883 | 12 |
val loop: unit -> unit |
| 6860 | 13 |
val sync_main: unit -> unit |
14 |
val sync_loop: unit -> unit |
|
| 5829 | 15 |
end; |
16 |
||
17 |
signature OUTER_SYNTAX = |
|
18 |
sig |
|
19 |
include BASIC_OUTER_SYNTAX |
|
| 6722 | 20 |
structure Keyword: |
21 |
sig |
|
22 |
val control: string |
|
23 |
val diag: string |
|
24 |
val thy_begin: string |
|
| 7104 | 25 |
val thy_switch: string |
| 6722 | 26 |
val thy_end: string |
27 |
val thy_heading: string |
|
28 |
val thy_decl: string |
|
| 9588 | 29 |
val thy_script: string |
| 6722 | 30 |
val thy_goal: string |
31 |
val qed: string |
|
| 6733 | 32 |
val qed_block: string |
| 8209 | 33 |
val qed_global: string |
| 9552 | 34 |
val prf_heading: string |
| 6722 | 35 |
val prf_goal: string |
36 |
val prf_block: string |
|
| 9056 | 37 |
val prf_open: string |
38 |
val prf_close: string |
|
| 6722 | 39 |
val prf_chain: string |
40 |
val prf_decl: string |
|
| 6868 | 41 |
val prf_asm: string |
| 7676 | 42 |
val prf_asm_goal: string |
| 6722 | 43 |
val prf_script: string |
44 |
val kinds: string list |
|
45 |
end |
|
| 5829 | 46 |
type token |
47 |
type parser |
|
| 6722 | 48 |
val command: string -> string -> string -> |
| 6373 | 49 |
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser |
| 9132 | 50 |
val markup_command: IsarOutput.markup -> string -> string -> string -> |
| 7789 | 51 |
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser |
| 6722 | 52 |
val improper_command: string -> string -> string -> |
| 6373 | 53 |
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser |
| 14687 | 54 |
val is_keyword: string -> bool |
| 7026 | 55 |
val dest_keywords: unit -> string list |
56 |
val dest_parsers: unit -> (string * string * string * bool) list |
|
| 5883 | 57 |
val print_outer_syntax: unit -> unit |
| 9223 | 58 |
val print_commands: Toplevel.transition -> Toplevel.transition |
| 5829 | 59 |
val add_keywords: string list -> unit |
60 |
val add_parsers: parser list -> unit |
|
| 12943 | 61 |
val check_text: string * Position.T -> bool -> Toplevel.state -> unit |
| 7940 | 62 |
val deps_thy: string -> bool -> Path.T -> string list * Path.T list |
| 6199 | 63 |
val load_thy: string -> bool -> bool -> Path.T -> unit |
| 14091 | 64 |
val isar: bool -> bool -> unit Toplevel.isar |
65 |
val isar_readstring: Position.T -> string -> (string list) Toplevel.isar |
|
| 5829 | 66 |
end; |
67 |
||
68 |
structure OuterSyntax: OUTER_SYNTAX = |
|
69 |
struct |
|
70 |
||
| 7750 | 71 |
structure T = OuterLex; |
| 6860 | 72 |
structure P = OuterParse; |
73 |
||
| 5829 | 74 |
|
75 |
(** outer syntax **) |
|
76 |
||
| 6722 | 77 |
(* command keyword classification *) |
78 |
||
79 |
structure Keyword = |
|
80 |
struct |
|
81 |
val control = "control"; |
|
82 |
val diag = "diag"; |
|
83 |
val thy_begin = "theory-begin"; |
|
| 7104 | 84 |
val thy_switch = "theory-switch"; |
| 6722 | 85 |
val thy_end = "theory-end"; |
86 |
val thy_heading = "theory-heading"; |
|
87 |
val thy_decl = "theory-decl"; |
|
| 9588 | 88 |
val thy_script = "theory-script"; |
| 6722 | 89 |
val thy_goal = "theory-goal"; |
90 |
val qed = "qed"; |
|
| 6733 | 91 |
val qed_block = "qed-block"; |
| 8209 | 92 |
val qed_global = "qed-global"; |
| 9552 | 93 |
val prf_heading = "proof-heading"; |
| 6722 | 94 |
val prf_goal = "proof-goal"; |
95 |
val prf_block = "proof-block"; |
|
| 9056 | 96 |
val prf_open = "proof-open"; |
97 |
val prf_close = "proof-close"; |
|
| 6722 | 98 |
val prf_chain = "proof-chain"; |
99 |
val prf_decl = "proof-decl"; |
|
| 6868 | 100 |
val prf_asm = "proof-asm"; |
| 7676 | 101 |
val prf_asm_goal = "proof-asm-goal"; |
| 6722 | 102 |
val prf_script = "proof-script"; |
103 |
||
| 9588 | 104 |
val kinds = [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, |
105 |
thy_goal, qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open, prf_close, |
|
106 |
prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script]; |
|
| 6722 | 107 |
end; |
108 |
||
109 |
||
| 5829 | 110 |
(* parsers *) |
111 |
||
| 7750 | 112 |
type token = T.token; |
| 5829 | 113 |
type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list; |
114 |
||
115 |
datatype parser = |
|
| 9132 | 116 |
Parser of string * (string * string * IsarOutput.markup option) * bool * parser_fn; |
| 5829 | 117 |
|
| 7750 | 118 |
fun parser int_only markup name comment kind parse = |
119 |
Parser (name, (comment, kind, markup), int_only, parse); |
|
| 5829 | 120 |
|
121 |
||
122 |
(* parse command *) |
|
123 |
||
| 6860 | 124 |
local |
| 6199 | 125 |
|
| 5829 | 126 |
fun command_body cmd (name, _) = |
| 7026 | 127 |
(case cmd name of |
128 |
Some (int_only, parse) => P.!!! (Scan.prompt (name ^ "# ") (parse >> pair int_only)) |
|
129 |
| None => sys_error ("no parser for outer syntax command " ^ quote name));
|
|
| 6860 | 130 |
|
131 |
fun terminator false = Scan.succeed () |
|
| 9132 | 132 |
| terminator true = P.group "end of input" (Scan.option P.sync -- P.semicolon >> K ()); |
| 6860 | 133 |
|
134 |
in |
|
| 5829 | 135 |
|
| 6860 | 136 |
fun command term cmd = |
| 9132 | 137 |
P.semicolon >> K None || |
| 6860 | 138 |
P.sync >> K None || |
| 7026 | 139 |
(P.position P.command :-- command_body cmd) --| terminator term |
| 6860 | 140 |
>> (fn ((name, pos), (int_only, f)) => |
141 |
Some (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |> |
|
142 |
Toplevel.interactive int_only |> f)); |
|
| 5829 | 143 |
|
| 6199 | 144 |
end; |
145 |
||
| 5829 | 146 |
|
147 |
||
| 9132 | 148 |
(** global outer syntax **) |
| 5829 | 149 |
|
| 7026 | 150 |
local |
151 |
||
152 |
val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon); |
|
| 7750 | 153 |
val global_parsers = |
| 9132 | 154 |
ref (Symtab.empty: (((string * string) * (bool * parser_fn)) * IsarOutput.markup option) |
155 |
Symtab.table); |
|
156 |
val global_markups = ref ([]: (string * IsarOutput.markup) list); |
|
| 5952 | 157 |
|
| 7026 | 158 |
fun change_lexicons f = |
159 |
let val lexs = f (! global_lexicons) in |
|
160 |
(case (op inter_string) (pairself Scan.dest_lexicon lexs) of |
|
161 |
[] => global_lexicons := lexs |
|
162 |
| bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads))
|
|
163 |
end; |
|
| 5829 | 164 |
|
| 7789 | 165 |
fun get_markup (ms, (name, (_, Some m))) = (name, m) :: ms |
166 |
| get_markup (ms, _) = ms; |
|
| 7750 | 167 |
|
168 |
fun make_markups () = global_markups := Symtab.foldl get_markup ([], ! global_parsers); |
|
| 9132 | 169 |
fun change_parsers f = (Library.change global_parsers f; make_markups ()); |
| 6722 | 170 |
|
| 7026 | 171 |
in |
172 |
||
| 7750 | 173 |
|
| 9132 | 174 |
(* access current syntax *) |
| 7026 | 175 |
|
176 |
(*Note: the syntax for files is statically determined at the very |
|
177 |
beginning; for interactive processing it may change dynamically.*) |
|
178 |
||
179 |
fun get_lexicons () = ! global_lexicons; |
|
180 |
fun get_parsers () = ! global_parsers; |
|
| 7750 | 181 |
fun get_parser () = apsome (#2 o #1) o curry Symtab.lookup (! global_parsers); |
| 7789 | 182 |
|
| 9132 | 183 |
fun is_markup kind name = |
184 |
(case assoc (! global_markups, name) of Some k => k = kind | None => false); |
|
185 |
fun markup kind = Scan.one (T.is_kind T.Command andf is_markup kind o T.val_of); |
|
| 5829 | 186 |
|
187 |
||
188 |
(* augment syntax *) |
|
189 |
||
| 7026 | 190 |
fun add_keywords keywords = change_lexicons (apfst (fn lex => |
191 |
(Scan.extend_lexicon lex (map Symbol.explode keywords)))); |
|
| 5829 | 192 |
|
| 7750 | 193 |
fun add_parser (tab, Parser (name, (comment, kind, markup), int_only, parse)) = |
| 5829 | 194 |
(if is_none (Symtab.lookup (tab, name)) then () |
195 |
else warning ("Redefined outer syntax command " ^ quote name);
|
|
| 7750 | 196 |
Symtab.update ((name, (((comment, kind), (int_only, parse)), markup)), tab)); |
| 5829 | 197 |
|
198 |
fun add_parsers parsers = |
|
| 7026 | 199 |
(change_parsers (fn tab => foldl add_parser (tab, parsers)); |
200 |
change_lexicons (apsnd (fn lex => Scan.extend_lexicon lex |
|
201 |
(map (fn Parser (name, _, _, _) => Symbol.explode name) parsers)))); |
|
202 |
||
203 |
end; |
|
| 5829 | 204 |
|
205 |
||
| 7026 | 206 |
(* print syntax *) |
207 |
||
| 14687 | 208 |
fun is_keyword s = Scan.is_literal (#1 (get_lexicons ())) (Symbol.explode s); |
| 7026 | 209 |
fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ())); |
210 |
||
211 |
fun dest_parsers () = |
|
| 7750 | 212 |
map (fn (name, (((cmt, kind), (int_only, _)), _)) => (name, cmt, kind, int_only)) |
| 7026 | 213 |
(Symtab.dest (get_parsers ())); |
| 5829 | 214 |
|
| 9223 | 215 |
fun print_outer_syntax () = |
| 7026 | 216 |
let |
217 |
fun pretty_cmd (name, comment, _, _) = |
|
218 |
Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
|
219 |
val (int_cmds, cmds) = partition #4 (dest_parsers ()); |
|
220 |
in |
|
| 8720 | 221 |
[Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())),
|
222 |
Pretty.big_list "proper commands:" (map pretty_cmd cmds), |
|
223 |
Pretty.big_list "improper commands (interactive-only):" (map pretty_cmd int_cmds)] |
|
| 9223 | 224 |
|> Pretty.chunks |> Pretty.writeln |
| 7026 | 225 |
end; |
| 5829 | 226 |
|
| 9223 | 227 |
val print_commands = Toplevel.imperative print_outer_syntax; |
| 7367 | 228 |
|
| 5829 | 229 |
|
230 |
||
| 9132 | 231 |
(** toplevel parsing **) |
| 5829 | 232 |
|
| 9132 | 233 |
(* basic sources *) |
| 6860 | 234 |
|
| 9132 | 235 |
fun toplevel_source term do_recover cmd src = |
236 |
let |
|
237 |
val no_terminator = |
|
238 |
Scan.unless P.semicolon (Scan.one (T.not_sync andf T.not_eof)); |
|
|
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
239 |
fun recover x = (Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [None]) x; |
| 9132 | 240 |
in |
241 |
src |
|
|
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
242 |
|> T.source_proper |
| 9132 | 243 |
|> Source.source T.stopper |
|
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
244 |
(Scan.bulk (P.$$$ "--" -- P.!!! P.text >> K None || P.not_eof >> Some)) |
|
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
245 |
(if do_recover then Some recover else None) |
|
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
246 |
|> Source.mapfilter I |
|
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
247 |
|> Source.source T.stopper (Scan.bulk (fn xs => P.!!! (command term (cmd ())) xs)) |
| 9132 | 248 |
(if do_recover then Some recover else None) |
249 |
|> Source.mapfilter I |
|
250 |
end; |
|
| 5829 | 251 |
|
| 7746 | 252 |
|
| 9132 | 253 |
(* interactive source of toplevel transformers *) |
| 5829 | 254 |
|
| 9132 | 255 |
fun isar term no_pos = |
256 |
Source.tty |
|
257 |
|> Symbol.source true |
|
258 |
|> T.source true get_lexicons |
|
259 |
(if no_pos then Position.none else Position.line_name 1 "stdin") |
|
260 |
|> toplevel_source term true get_parser; |
|
| 6199 | 261 |
|
262 |
||
| 14091 | 263 |
(* string source of transformers (for Proof General) *) |
264 |
||
265 |
fun isar_readstring pos str = |
|
266 |
Source.of_string str |
|
267 |
|> Symbol.source false |
|
268 |
|> T.source false get_lexicons pos |
|
269 |
|> toplevel_source false true get_parser; |
|
270 |
||
271 |
||
| 7746 | 272 |
|
| 9132 | 273 |
(** read theory **) |
| 6247 | 274 |
|
| 12943 | 275 |
(* check_text *) |
276 |
||
277 |
fun check_text s true state = (IsarOutput.eval_antiquote (#1 (get_lexicons ())) state s; ()) |
|
278 |
| check_text _ false _ = (); |
|
279 |
||
280 |
||
| 9132 | 281 |
(* deps_thy *) |
| 6199 | 282 |
|
| 7940 | 283 |
fun deps_thy name ml path = |
| 6199 | 284 |
let |
| 7735 | 285 |
val src = Source.of_string (File.read path); |
286 |
val pos = Path.position path; |
|
| 9132 | 287 |
val (name', parents, files) = ThyHeader.scan (src, pos); |
| 6199 | 288 |
val ml_path = ThyLoad.ml_path name; |
| 7940 | 289 |
val ml_file = if ml andalso is_some (ThyLoad.check_file ml_path) then [ml_path] else []; |
| 6199 | 290 |
in |
291 |
if name <> name' then |
|
| 7940 | 292 |
error ("Filename " ^ quote (Path.pack path) ^
|
| 8078 | 293 |
" does not agree with theory name " ^ quote name') |
| 6247 | 294 |
else (parents, map (Path.unpack o #1) files @ ml_file) |
| 6199 | 295 |
end; |
296 |
||
| 7746 | 297 |
|
| 9132 | 298 |
(* load_thy *) |
| 6199 | 299 |
|
| 7746 | 300 |
local |
301 |
||
| 7940 | 302 |
fun try_ml_file name time = |
| 6199 | 303 |
let |
304 |
val path = ThyLoad.ml_path name; |
|
| 7940 | 305 |
val tr = Toplevel.imperative (fn () => ThyInfo.load_file time path); |
| 6247 | 306 |
val tr_name = if time then "time_use" else "use"; |
| 6199 | 307 |
in |
| 7243 | 308 |
if is_none (ThyLoad.check_file path) then () |
| 9132 | 309 |
else Toplevel.excursion [Toplevel.empty |> Toplevel.name tr_name |> tr] |
| 6199 | 310 |
end; |
311 |
||
| 9132 | 312 |
fun parse_thy src = |
313 |
src |
|
314 |
|> toplevel_source false false (K (get_parser ())) |
|
| 7746 | 315 |
|> Source.exhaust; |
| 5829 | 316 |
|
| 6247 | 317 |
fun run_thy name path = |
| 7683 | 318 |
let |
| 7735 | 319 |
val pos = Path.position path; |
| 9132 | 320 |
val text = Library.untabify (explode (File.read path)); |
321 |
val text_src = Source.of_list text; |
|
322 |
fun present_text () = Source.exhaust (Symbol.source false text_src); |
|
| 7683 | 323 |
in |
| 7735 | 324 |
Present.init_theory name; |
| 9132 | 325 |
Present.verbatim_source name present_text; |
326 |
if ThyHeader.is_old (text_src, pos) then (ThySyn.load_thy name text; |
|
327 |
Present.old_symbol_source name present_text) (*note: text presented twice*) |
|
328 |
else |
|
329 |
let |
|
|
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
330 |
val tok_src = text_src |
|
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
331 |
|> Symbol.source false |
|
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
332 |
|> T.source false (K (get_lexicons ())) pos |
|
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
10749
diff
changeset
|
333 |
|> Source.exhausted; |
| 9132 | 334 |
val out = Toplevel.excursion_result |
335 |
(IsarOutput.parse_thy markup (#1 (get_lexicons ())) |
|
336 |
(parse_thy tok_src) tok_src); |
|
337 |
in Present.theory_output name (Buffer.content out) end |
|
| 6247 | 338 |
end; |
| 6199 | 339 |
|
| 7746 | 340 |
in |
341 |
||
| 6199 | 342 |
fun load_thy name ml time path = |
| 6247 | 343 |
(if time then |
344 |
timeit (fn () => |
|
345 |
(writeln ("\n**** Starting theory " ^ quote name ^ " ****");
|
|
| 9036 | 346 |
run_thy name path; |
| 6247 | 347 |
writeln ("**** Finished theory " ^ quote name ^ " ****\n")))
|
348 |
else run_thy name path; |
|
349 |
Context.context (ThyInfo.get_theory name); |
|
| 7940 | 350 |
if ml then try_ml_file name time else ()); |
| 5829 | 351 |
|
| 7746 | 352 |
end; |
353 |
||
| 5829 | 354 |
|
355 |
||
356 |
(** the read-eval-print loop **) |
|
357 |
||
| 5923 | 358 |
(* main loop *) |
359 |
||
| 7333 | 360 |
fun gen_loop term no_pos = |
361 |
(Context.reset_context (); |
|
362 |
Toplevel.loop (isar term no_pos)); |
|
| 5829 | 363 |
|
| 7333 | 364 |
fun gen_main term no_pos = |
| 5829 | 365 |
(Toplevel.set_state Toplevel.toplevel; |
| 6199 | 366 |
writeln (Session.welcome ()); |
| 7333 | 367 |
gen_loop term no_pos); |
| 6860 | 368 |
|
| 7333 | 369 |
fun main () = gen_main false false; |
370 |
fun loop () = gen_loop false false; |
|
371 |
fun sync_main () = gen_main true true; |
|
372 |
fun sync_loop () = gen_loop true true; |
|
| 5829 | 373 |
|
374 |
||
375 |
(* help *) |
|
376 |
||
377 |
fun help () = |
|
378 |
writeln ("This is Isabelle's underlying ML system (" ^ ml_system ^ ");\n\
|
|
| 9223 | 379 |
\invoke 'Isar.loop();' to get back to the Isar read-eval-print loop."); |
| 5829 | 380 |
|
381 |
||
| 6373 | 382 |
(*final declarations of this structure!*) |
| 7789 | 383 |
val command = parser false None; |
| 9132 | 384 |
val markup_command = parser false o Some; |
| 7789 | 385 |
val improper_command = parser true None; |
| 6685 | 386 |
|
| 6373 | 387 |
|
| 5829 | 388 |
end; |
389 |
||
| 6199 | 390 |
(*setup theory syntax dependent operations*) |
391 |
ThyLoad.deps_thy_fn := OuterSyntax.deps_thy; |
|
392 |
ThyLoad.load_thy_fn := OuterSyntax.load_thy; |
|
393 |
structure ThyLoad: THY_LOAD = ThyLoad; |
|
394 |
||
| 5829 | 395 |
structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax; |
396 |
open BasicOuterSyntax; |