1 (* Title: Pure/Isar/outer_syntax.ML
3 Author: Markus Wenzel, TU Muenchen
5 The global Isabelle/Isar outer syntax.
8 signature BASIC_OUTER_SYNTAX =
10 val main: unit -> unit
11 val loop: unit -> unit
12 val sync_main: unit -> unit
13 val sync_loop: unit -> unit
14 val help: unit -> unit
17 signature OUTER_SYNTAX =
19 include BASIC_OUTER_SYNTAX
25 val thy_switch: string
27 val thy_heading: string
32 val qed_global: string
38 val prf_asm_goal: string
39 val prf_script: string
40 val kinds: string list
44 val command: string -> string -> string ->
45 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
46 val markup_command: string -> string -> string ->
47 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
48 val verbatim_command: string -> string -> string ->
49 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
50 val improper_command: string -> string -> string ->
51 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
52 val dest_keywords: unit -> string list
53 val dest_parsers: unit -> (string * string * string * bool) list
54 val print_outer_syntax: unit -> unit
55 val print_help: Toplevel.transition -> Toplevel.transition
56 val add_keywords: string list -> unit
57 val add_parsers: parser list -> unit
58 val theory_header: token list -> (string * string list * (string * bool) list) * token list
59 val deps_thy: string -> bool -> Path.T -> string list * Path.T list
60 val load_thy: string -> bool -> bool -> Path.T -> unit
61 val isar: bool -> bool -> Toplevel.isar
64 structure OuterSyntax: OUTER_SYNTAX =
67 structure T = OuterLex;
68 structure P = OuterParse;
73 (* command keyword classification *)
77 val control = "control";
79 val thy_begin = "theory-begin";
80 val thy_switch = "theory-switch";
81 val thy_end = "theory-end";
82 val thy_heading = "theory-heading";
83 val thy_decl = "theory-decl";
84 val thy_goal = "theory-goal";
86 val qed_block = "qed-block";
87 val qed_global = "qed-global";
88 val prf_goal = "proof-goal";
89 val prf_block = "proof-block";
90 val prf_chain = "proof-chain";
91 val prf_decl = "proof-decl";
92 val prf_asm = "proof-asm";
93 val prf_asm_goal = "proof-asm-goal";
94 val prf_script = "proof-script";
96 val kinds = [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_goal,
97 qed, qed_block, qed_global, prf_goal, prf_block, prf_chain, prf_decl, prf_asm, prf_asm_goal,
104 type token = T.token;
105 type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list;
108 Parser of string * (string * string * bool option) * bool * parser_fn;
110 fun parser int_only markup name comment kind parse =
111 Parser (name, (comment, kind, markup), int_only, parse);
118 fun command_body cmd (name, _) =
120 Some (int_only, parse) => P.!!! (Scan.prompt (name ^ "# ") (parse >> pair int_only))
121 | None => sys_error ("no parser for outer syntax command " ^ quote name));
123 fun terminator false = Scan.succeed ()
124 | terminator true = P.group "terminator" (Scan.option P.sync -- P.$$$ ";" >> K ());
128 fun command term cmd =
129 P.$$$ ";" >> K None ||
131 (P.position P.command :-- command_body cmd) --| terminator term
132 >> (fn ((name, pos), (int_only, f)) =>
133 Some (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
134 Toplevel.interactive int_only |> f));
140 (** global syntax state **)
144 val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon);
146 ref (Symtab.empty: (((string * string) * (bool * parser_fn)) * bool option) Symtab.table);
147 val global_markups = ref ([]: (string * bool) list);
149 fun change_lexicons f =
150 let val lexs = f (! global_lexicons) in
151 (case (op inter_string) (pairself Scan.dest_lexicon lexs) of
152 [] => global_lexicons := lexs
153 | bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads))
156 fun get_markup (ms, (name, (_, Some m))) = (name, m) :: ms
157 | get_markup (ms, _) = ms;
159 fun make_markups () = global_markups := Symtab.foldl get_markup ([], ! global_parsers);
160 fun change_parsers f = (global_parsers := f (! global_parsers); make_markups ());
165 (* get current syntax *)
167 (*Note: the syntax for files is statically determined at the very
168 beginning; for interactive processing it may change dynamically.*)
170 fun get_lexicons () = ! global_lexicons;
171 fun get_parsers () = ! global_parsers;
172 fun get_parser () = apsome (#2 o #1) o curry Symtab.lookup (! global_parsers);
174 fun lookup_markup name = assoc (! global_markups, name);
175 fun is_markup name = if_none (lookup_markup name) false;
176 fun is_verbatim name = if_none (apsome not (lookup_markup name)) false;
181 fun add_keywords keywords = change_lexicons (apfst (fn lex =>
182 (Scan.extend_lexicon lex (map Symbol.explode keywords))));
184 fun add_parser (tab, Parser (name, (comment, kind, markup), int_only, parse)) =
185 (if is_none (Symtab.lookup (tab, name)) then ()
186 else warning ("Redefined outer syntax command " ^ quote name);
187 Symtab.update ((name, (((comment, kind), (int_only, parse)), markup)), tab));
189 fun add_parsers parsers =
190 (change_parsers (fn tab => foldl add_parser (tab, parsers));
191 change_lexicons (apsnd (fn lex => Scan.extend_lexicon lex
192 (map (fn Parser (name, _, _, _) => Symbol.explode name) parsers))));
199 fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ()));
201 fun dest_parsers () =
202 map (fn (name, (((cmt, kind), (int_only, _)), _)) => (name, cmt, kind, int_only))
203 (Symtab.dest (get_parsers ()));
205 fun print_outer_syntax () =
207 fun pretty_cmd (name, comment, _, _) =
208 Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
209 val (int_cmds, cmds) = partition #4 (dest_parsers ());
211 Pretty.writeln (Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())));
212 Pretty.writeln (Pretty.big_list "proper commands:" (map pretty_cmd cmds));
213 Pretty.writeln (Pretty.big_list "improper commands (interactive-only):"
214 (map pretty_cmd int_cmds))
218 Toplevel.keep (fn state =>
219 let val opt_thy = try Toplevel.theory_of state in
220 print_outer_syntax ();
221 Method.help_methods opt_thy;
222 Attrib.help_attributes opt_thy
229 (* special keywords *)
231 val headerN = "header";
232 val theoryN = "theory";
234 val theory_keyword = P.$$$ theoryN;
235 val header_keyword = P.$$$ headerN;
236 val semicolon = P.$$$ ";";
244 Scan.unless semicolon (Scan.one (T.not_sync andf T.not_eof));
246 val recover = Scan.prompt "recover# " (Scan.repeat no_terminator);
250 fun source term do_recover cmd src =
252 |> Source.source T.stopper
253 (Scan.bulk (fn xs => P.!!! (command term (cmd ())) xs))
254 (if do_recover then Some recover else None)
255 |> Source.mapfilter I;
259 fun token_source (src, pos) =
261 |> Symbol.source false
262 |> T.source false (K (get_lexicons ())) pos;
264 fun filter_proper src =
266 |> Source.filter T.is_proper;
271 fun scan_header get_lex scan (src, pos) =
273 |> Symbol.source false
274 |> T.source false (fn () => (get_lex (), Scan.empty_lexicon)) pos
276 |> Source.source T.stopper (Scan.single scan) None
277 |> (fst o the o Source.get_single);
280 (* detect new/old header *)
284 val check_header_lexicon = Scan.make_lexicon [Symbol.explode headerN, Symbol.explode theoryN];
285 val check_header = Scan.option (header_keyword || theory_keyword);
289 fun is_old_theory src = is_none (scan_header (K check_header_lexicon) check_header src);
294 (* deps_thy --- inspect theory header *)
299 Scan.make_lexicon (map Symbol.explode ["(", ")", "+", ":", ";", "=", "files", headerN, theoryN]);
302 (P.$$$ "(" |-- P.!!! (P.name --| P.$$$ ")")) >> rpair false || P.name >> rpair true;
307 (P.name -- (P.$$$ "=" |-- P.enum1 "+" P.name) --
308 Scan.optional (P.$$$ "files" |-- P.!!! (Scan.repeat1 file_name)) [] --| P.$$$ ":")
309 >> (fn ((A, Bs), files) => (A, Bs, files));
312 header_keyword |-- (P.!!! (P.text -- Scan.option semicolon -- theory_keyword |-- theory_header))
313 || theory_keyword |-- P.!!! theory_header;
316 P.!!! (P.group "theory header"
317 (P.name -- (P.$$$ "=" |-- P.name -- Scan.repeat (P.$$$ "+" |-- P.name))))
318 >> (fn (A, (B, Bs)) => (A, B :: Bs, []: (string * bool) list));
320 fun deps_thy name ml path =
322 val src = Source.of_string (File.read path);
323 val pos = Path.position path;
324 val (name', parents, files) =
325 (*unfortunately, old-style headers dynamically depend on the current lexicon*)
326 if is_old_theory (src, pos) then
327 scan_header ThySyn.get_lexicon (Scan.error old_header) (src, pos)
328 else scan_header (K header_lexicon) (Scan.error new_header) (src, pos);
330 val ml_path = ThyLoad.ml_path name;
331 val ml_file = if ml andalso is_some (ThyLoad.check_file ml_path) then [ml_path] else [];
333 if name <> name' then
334 error ("Filename " ^ quote (Path.pack path) ^
335 " does not agree with theory name " ^ quote name')
336 else (parents, map (Path.unpack o #1) files @ ml_file)
342 (* present theory source *)
346 val indent_prop = Scan.one T.is_indent -- Scan.one T.is_proper;
347 val improp = Scan.unless indent_prop (Scan.one (not o T.is_proper));
348 val improper_keep_indent = Scan.repeat improp;
350 val improper = Scan.any (not o T.is_proper);
353 (improper -- semicolon) |-- improper_keep_indent ||
354 improper_keep_indent;
356 val markup = Scan.one (T.is_kind T.Command andf is_markup o T.val_of) >> T.val_of;
357 val verbatim = Scan.one (T.is_kind T.Command andf is_verbatim o T.val_of);
360 improper |-- markup -- (improper |-- P.text --| improper_end) >> Present.markup_token ||
361 (P.$$$ "--" >> K "cmt") -- (improper |-- P.text) >> Present.markup_token ||
362 (improper -- verbatim -- improper) |-- P.text --| improper_end >> Present.verbatim_token ||
363 Scan.one T.not_eof >> Present.basic_token;
367 (*note: lazy evaluation ahead*)
369 fun present_toks text pos () =
370 token_source (Source.of_list (Library.untabify text), pos)
371 |> Source.source T.stopper (Scan.bulk present_token) None
374 fun present_text text () =
375 Source.exhaust (Symbol.source false (Source.of_list (Library.untabify text)));
380 (* load_thy --- read text (including header) *)
384 fun try_ml_file name time =
386 val path = ThyLoad.ml_path name;
387 val tr = Toplevel.imperative (fn () => ThyInfo.load_file time path);
388 val tr_name = if time then "time_use" else "use";
390 if is_none (ThyLoad.check_file path) then ()
391 else Toplevel.excursion_error [Toplevel.empty |> Toplevel.name tr_name |> tr]
394 fun parse_thy src_pos =
398 |> source false false (K (get_parser ()))
401 fun run_thy name path =
403 val text = explode (File.read path);
404 val src = Source.of_list text;
405 val pos = Path.position path;
407 Present.init_theory name;
408 Present.verbatim_source name (present_text text);
409 if is_old_theory (src, pos) then (ThySyn.load_thy name text;
410 Present.old_symbol_source name (present_text text)) (*note: text presented twice!*)
411 else (Toplevel.excursion_error (parse_thy (src, pos));
412 Present.token_source name (present_toks text pos))
417 fun load_thy name ml time path =
420 (writeln ("\n**** Starting theory " ^ quote name ^ " ****");
421 setmp Goals.proof_timing true (run_thy name) path;
422 writeln ("**** Finished theory " ^ quote name ^ " ****\n")))
423 else run_thy name path;
424 Context.context (ThyInfo.get_theory name);
425 if ml then try_ml_file name time else ());
430 (* interactive source of state transformers *)
432 fun isar term no_pos =
434 |> Symbol.source true
435 |> T.source true get_lexicons
436 (if no_pos then Position.none else Position.line_name 1 "stdin")
438 |> source term true get_parser;
442 (** the read-eval-print loop **)
446 fun gen_loop term no_pos =
447 (Context.reset_context ();
448 Toplevel.loop (isar term no_pos));
450 fun gen_main term no_pos =
451 (Toplevel.set_state Toplevel.toplevel;
452 writeln (Session.welcome ());
453 gen_loop term no_pos);
455 fun main () = gen_main false false;
456 fun loop () = gen_loop false false;
457 fun sync_main () = gen_main true true;
458 fun sync_loop () = gen_loop true true;
464 writeln ("This is Isabelle's underlying ML system (" ^ ml_system ^ ");\n\
465 \invoke 'loop();' to enter the Isar loop.");
468 (*final declarations of this structure!*)
469 val command = parser false None;
470 val markup_command = parser false (Some true);
471 val verbatim_command = parser false (Some false);
472 val improper_command = parser true None;
477 (*setup theory syntax dependent operations*)
478 ThyLoad.deps_thy_fn := OuterSyntax.deps_thy;
479 ThyLoad.load_thy_fn := OuterSyntax.load_thy;
480 structure ThyLoad: THY_LOAD = ThyLoad;
482 structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax;
483 open BasicOuterSyntax;