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
37 val prf_asm_goal: string
38 val prf_script: string
39 val kinds: string list
43 val command: string -> string -> string ->
44 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
45 val markup_command: string -> string -> string ->
46 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
47 val verbatim_command: string -> string -> string ->
48 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
49 val improper_command: string -> string -> string ->
50 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
51 val dest_keywords: unit -> string list
52 val dest_parsers: unit -> (string * string * string * bool) list
53 val print_outer_syntax: unit -> unit
54 val print_help: Toplevel.transition -> Toplevel.transition
55 val add_keywords: string list -> unit
56 val add_parsers: parser list -> unit
57 val theory_header: token list -> (string * string list * (string * bool) list) * token list
58 val deps_thy: string -> Path.T -> string list * Path.T list
59 val load_thy: string -> bool -> bool -> Path.T -> unit
60 val isar: bool -> bool -> Toplevel.isar
63 structure OuterSyntax: OUTER_SYNTAX =
66 structure T = OuterLex;
67 structure P = OuterParse;
72 (* command keyword classification *)
76 val control = "control";
78 val thy_begin = "theory-begin";
79 val thy_switch = "theory-switch";
80 val thy_end = "theory-end";
81 val thy_heading = "theory-heading";
82 val thy_decl = "theory-decl";
83 val thy_goal = "theory-goal";
85 val qed_block = "qed-block";
86 val prf_goal = "proof-goal";
87 val prf_block = "proof-block";
88 val prf_chain = "proof-chain";
89 val prf_decl = "proof-decl";
90 val prf_asm = "proof-asm";
91 val prf_asm_goal = "proof-asm-goal";
92 val prf_script = "proof-script";
94 val kinds = [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_goal,
95 qed, qed_block, prf_goal, prf_block, prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script];
101 type token = T.token;
102 type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list;
105 Parser of string * (string * string * bool option) * bool * parser_fn;
107 fun parser int_only markup name comment kind parse =
108 Parser (name, (comment, kind, markup), int_only, parse);
115 fun command_body cmd (name, _) =
117 Some (int_only, parse) => P.!!! (Scan.prompt (name ^ "# ") (parse >> pair int_only))
118 | None => sys_error ("no parser for outer syntax command " ^ quote name));
120 fun terminator false = Scan.succeed ()
121 | terminator true = P.group "terminator" (Scan.option P.sync -- P.$$$ ";" >> K ());
125 fun command term cmd =
126 P.$$$ ";" >> K None ||
128 (P.position P.command :-- command_body cmd) --| terminator term
129 >> (fn ((name, pos), (int_only, f)) =>
130 Some (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
131 Toplevel.interactive int_only |> f));
137 (** global syntax state **)
141 val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon);
143 ref (Symtab.empty: (((string * string) * (bool * parser_fn)) * bool option) Symtab.table);
144 val global_markups = ref ([]: (string * bool) list);
146 fun change_lexicons f =
147 let val lexs = f (! global_lexicons) in
148 (case (op inter_string) (pairself Scan.dest_lexicon lexs) of
149 [] => global_lexicons := lexs
150 | bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads))
153 fun get_markup (ms, (name, (_, Some m))) = (name, m) :: ms
154 | get_markup (ms, _) = ms;
156 fun make_markups () = global_markups := Symtab.foldl get_markup ([], ! global_parsers);
157 fun change_parsers f = (global_parsers := f (! global_parsers); make_markups ());
162 (* get current syntax *)
164 (*Note: the syntax for files is statically determined at the very
165 beginning; for interactive processing it may change dynamically.*)
167 fun get_lexicons () = ! global_lexicons;
168 fun get_parsers () = ! global_parsers;
169 fun get_parser () = apsome (#2 o #1) o curry Symtab.lookup (! global_parsers);
171 fun lookup_markup name = assoc (! global_markups, name);
172 fun is_markup name = if_none (lookup_markup name) false;
173 fun is_verbatim name = if_none (apsome not (lookup_markup name)) false;
178 fun add_keywords keywords = change_lexicons (apfst (fn lex =>
179 (Scan.extend_lexicon lex (map Symbol.explode keywords))));
181 fun add_parser (tab, Parser (name, (comment, kind, markup), int_only, parse)) =
182 (if is_none (Symtab.lookup (tab, name)) then ()
183 else warning ("Redefined outer syntax command " ^ quote name);
184 Symtab.update ((name, (((comment, kind), (int_only, parse)), markup)), tab));
186 fun add_parsers parsers =
187 (change_parsers (fn tab => foldl add_parser (tab, parsers));
188 change_lexicons (apsnd (fn lex => Scan.extend_lexicon lex
189 (map (fn Parser (name, _, _, _) => Symbol.explode name) parsers))));
196 fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ()));
198 fun dest_parsers () =
199 map (fn (name, (((cmt, kind), (int_only, _)), _)) => (name, cmt, kind, int_only))
200 (Symtab.dest (get_parsers ()));
202 fun print_outer_syntax () =
204 fun pretty_cmd (name, comment, _, _) =
205 Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
206 val (int_cmds, cmds) = partition #4 (dest_parsers ());
208 Pretty.writeln (Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())));
209 Pretty.writeln (Pretty.big_list "proper commands:" (map pretty_cmd cmds));
210 Pretty.writeln (Pretty.big_list "improper commands (interactive-only):"
211 (map pretty_cmd int_cmds))
215 Toplevel.keep (fn state =>
216 let val opt_thy = try Toplevel.theory_of state in
217 print_outer_syntax ();
218 Method.help_methods opt_thy;
219 Attrib.help_attributes opt_thy
226 (* special keywords *)
228 val headerN = "header";
229 val theoryN = "theory";
231 val theory_keyword = P.$$$ theoryN;
232 val header_keyword = P.$$$ headerN;
233 val semicolon = P.$$$ ";";
241 Scan.unless semicolon (Scan.one (T.not_sync andf T.not_eof));
243 val recover = Scan.prompt "recover# " (Scan.repeat no_terminator);
247 fun source term do_recover cmd src =
249 |> Source.source T.stopper
250 (Scan.bulk (fn xs => P.!!! (command term (cmd ())) xs))
251 (if do_recover then Some recover else None)
252 |> Source.mapfilter I;
256 fun token_source (src, pos) =
258 |> Symbol.source false
259 |> T.source false (K (get_lexicons ())) pos;
261 fun filter_proper src =
263 |> Source.filter T.is_proper;
268 fun scan_header get_lex scan (src, pos) =
270 |> Symbol.source false
271 |> T.source false (fn () => (get_lex (), Scan.empty_lexicon)) pos
273 |> Source.source T.stopper (Scan.single scan) None
274 |> (fst o the o Source.get_single);
277 (* detect new/old header *)
281 val check_header_lexicon = Scan.make_lexicon [Symbol.explode headerN, Symbol.explode theoryN];
282 val check_header = Scan.option (header_keyword || theory_keyword);
286 fun is_old_theory src = is_none (scan_header (K check_header_lexicon) check_header src);
291 (* deps_thy --- inspect theory header *)
296 Scan.make_lexicon (map Symbol.explode ["(", ")", "+", ":", ";", "=", "files", headerN, theoryN]);
299 (P.$$$ "(" |-- P.!!! (P.name --| P.$$$ ")")) >> rpair false || P.name >> rpair true;
304 (P.name -- (P.$$$ "=" |-- P.enum1 "+" P.name) --
305 Scan.optional (P.$$$ "files" |-- P.!!! (Scan.repeat1 file_name)) [] --| P.$$$ ":")
306 >> (fn ((A, Bs), files) => (A, Bs, files));
309 header_keyword |-- (P.!!! (P.text -- Scan.option semicolon -- theory_keyword |-- theory_header))
310 || theory_keyword |-- P.!!! theory_header;
313 P.!!! (P.group "theory header"
314 (P.name -- (P.$$$ "=" |-- P.name -- Scan.repeat (P.$$$ "+" |-- P.name))))
315 >> (fn (A, (B, Bs)) => (A, B :: Bs, []: (string * bool) list));
317 fun deps_thy name path =
319 val src = Source.of_string (File.read path);
320 val pos = Path.position path;
321 val (name', parents, files) =
322 (*Note: old style headers dynamically depend on the current lexicon :-( *)
323 if is_old_theory (src, pos) then
324 scan_header ThySyn.get_lexicon (Scan.error old_header) (src, pos)
325 else scan_header (K header_lexicon) (Scan.error new_header) (src, pos);
327 val ml_path = ThyLoad.ml_path name;
328 val ml_file = if is_none (ThyLoad.check_file ml_path) then [] else [ml_path];
330 if name <> name' then
331 error ("Filename " ^ quote (Path.pack path) ^ " does not match theory name " ^ quote name)
332 else (parents, map (Path.unpack o #1) files @ ml_file)
338 (* present theory source *)
342 val improper = Scan.any (not o T.is_proper);
343 val markup = Scan.one (T.is_kind T.Command andf is_markup o T.val_of) >> T.val_of;
344 val verbatim = Scan.one (T.is_kind T.Command andf is_verbatim o T.val_of);
347 improper |-- markup -- (improper |-- P.text --| (improper -- Scan.option semicolon -- improper))
348 >> Present.markup_token ||
349 (P.$$$ "--" >> K "cmt") -- (improper |-- P.text) >> Present.markup_token ||
350 (improper -- verbatim -- improper) |-- P.text --| (improper -- Scan.option semicolon -- improper)
351 >> Present.verbatim_token ||
352 Scan.one T.not_eof >> Present.basic_token;
356 (*note: lazy evaluation ahead*)
358 fun present_toks text pos () =
359 token_source (Source.of_list (Library.untabify text), pos)
360 |> Source.source T.stopper (Scan.bulk present_token) None
363 fun present_text text () =
364 Source.exhaust (Symbol.source false (Source.of_list (Library.untabify text)));
369 (* load_thy --- read text (including header) *)
373 fun try_ml_file name ml time =
375 val path = ThyLoad.ml_path name;
376 val tr = Toplevel.imperative (fn () => ThyInfo.may_load_file ml time path);
377 val tr_name = if time then "time_use" else "use";
379 if is_none (ThyLoad.check_file path) then ()
380 else Toplevel.excursion_error [Toplevel.empty |> Toplevel.name tr_name |> tr]
383 fun parse_thy src_pos =
387 |> source false false (K (get_parser ()))
390 fun run_thy name path =
392 val text = explode (File.read path);
393 val src = Source.of_list text;
394 val pos = Path.position path;
396 Present.init_theory name;
397 Present.verbatim_source name (present_text text);
398 if is_old_theory (src, pos) then ThySyn.load_thy name text
399 else (Toplevel.excursion_error (parse_thy (src, pos));
400 Present.token_source name (present_toks text pos))
405 fun load_thy name ml time path =
408 (writeln ("\n**** Starting theory " ^ quote name ^ " ****");
409 setmp Goals.proof_timing true (run_thy name) path;
410 writeln ("**** Finished theory " ^ quote name ^ " ****\n")))
411 else run_thy name path;
412 Context.context (ThyInfo.get_theory name);
413 try_ml_file name ml time);
418 (* interactive source of state transformers *)
420 fun isar term no_pos =
422 |> Symbol.source true
423 |> T.source true get_lexicons
424 (if no_pos then Position.none else Position.line_name 1 "stdin")
426 |> source term true get_parser;
430 (** the read-eval-print loop **)
434 fun gen_loop term no_pos =
435 (Context.reset_context ();
436 Toplevel.loop (isar term no_pos));
438 fun gen_main term no_pos =
439 (Toplevel.set_state Toplevel.toplevel;
440 ml_prompts "ML> " "ML# ";
441 writeln (Session.welcome ());
442 gen_loop term no_pos);
444 fun main () = gen_main false false;
445 fun loop () = gen_loop false false;
446 fun sync_main () = gen_main true true;
447 fun sync_loop () = gen_loop true true;
453 writeln ("This is Isabelle's underlying ML system (" ^ ml_system ^ ");\n\
454 \invoke 'loop();' to enter the Isar loop.");
457 (*final declarations of this structure!*)
458 val command = parser false None;
459 val markup_command = parser false (Some true);
460 val verbatim_command = parser false (Some false);
461 val improper_command = parser true None;
466 (*setup theory syntax dependent operations*)
467 ThyLoad.deps_thy_fn := OuterSyntax.deps_thy;
468 ThyLoad.load_thy_fn := OuterSyntax.load_thy;
469 structure ThyLoad: THY_LOAD = ThyLoad;
471 structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax;
472 open BasicOuterSyntax;