1 (* Title: Pure/Isar/outer_syntax.ML
3 Author: Markus Wenzel, TU Muenchen
4 License: GPL (GNU GENERAL PUBLIC LICENSE)
6 The global Isabelle/Isar outer syntax.
9 signature BASIC_OUTER_SYNTAX =
11 val main: unit -> unit
12 val loop: unit -> unit
13 val sync_main: unit -> unit
14 val sync_loop: unit -> unit
15 val help: unit -> unit
18 signature OUTER_SYNTAX =
20 include BASIC_OUTER_SYNTAX
26 val thy_switch: string
28 val thy_heading: string
33 val qed_global: string
39 val prf_asm_goal: string
40 val prf_script: string
41 val kinds: string list
45 val command: string -> string -> string ->
46 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
47 val markup_command: string -> string -> string ->
48 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
49 val markup_env_command: string -> string -> string ->
50 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
51 val verbatim_command: string -> string -> string ->
52 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
53 val improper_command: string -> string -> string ->
54 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
55 val dest_keywords: unit -> string list
56 val dest_parsers: unit -> (string * string * string * bool) list
57 val print_outer_syntax: unit -> unit
58 val print_help: Toplevel.transition -> Toplevel.transition
59 val add_keywords: string list -> unit
60 val add_parsers: parser list -> unit
61 val theory_header: token list -> (string * string list * (string * bool) list) * token list
62 val deps_thy: string -> bool -> Path.T -> string list * Path.T list
63 val load_thy: string -> bool -> bool -> Path.T -> unit
64 val isar: bool -> bool -> Toplevel.isar
67 structure OuterSyntax: OUTER_SYNTAX =
70 structure T = OuterLex;
71 structure P = OuterParse;
76 (* command keyword classification *)
80 val control = "control";
82 val thy_begin = "theory-begin";
83 val thy_switch = "theory-switch";
84 val thy_end = "theory-end";
85 val thy_heading = "theory-heading";
86 val thy_decl = "theory-decl";
87 val thy_goal = "theory-goal";
89 val qed_block = "qed-block";
90 val qed_global = "qed-global";
91 val prf_goal = "proof-goal";
92 val prf_block = "proof-block";
93 val prf_chain = "proof-chain";
94 val prf_decl = "proof-decl";
95 val prf_asm = "proof-asm";
96 val prf_asm_goal = "proof-asm-goal";
97 val prf_script = "proof-script";
99 val kinds = [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_goal,
100 qed, qed_block, qed_global, prf_goal, prf_block, prf_chain, prf_decl, prf_asm, prf_asm_goal,
107 type token = T.token;
108 type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list;
110 datatype markup_kind = Markup | MarkupEnv | Verbatim;
113 Parser of string * (string * string * markup_kind option) * bool * parser_fn;
115 fun parser int_only markup name comment kind parse =
116 Parser (name, (comment, kind, markup), int_only, parse);
123 fun command_body cmd (name, _) =
125 Some (int_only, parse) => P.!!! (Scan.prompt (name ^ "# ") (parse >> pair int_only))
126 | None => sys_error ("no parser for outer syntax command " ^ quote name));
128 fun terminator false = Scan.succeed ()
129 | terminator true = P.group "terminator" (Scan.option P.sync -- P.$$$ ";" >> K ());
133 fun command term cmd =
134 P.$$$ ";" >> K None ||
136 (P.position P.command :-- command_body cmd) --| terminator term
137 >> (fn ((name, pos), (int_only, f)) =>
138 Some (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
139 Toplevel.interactive int_only |> f));
145 (** global syntax state **)
149 val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon);
151 ref (Symtab.empty: (((string * string) * (bool * parser_fn)) * markup_kind option) Symtab.table);
152 val global_markups = ref ([]: (string * markup_kind) list);
154 fun change_lexicons f =
155 let val lexs = f (! global_lexicons) in
156 (case (op inter_string) (pairself Scan.dest_lexicon lexs) of
157 [] => global_lexicons := lexs
158 | bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads))
161 fun get_markup (ms, (name, (_, Some m))) = (name, m) :: ms
162 | get_markup (ms, _) = ms;
164 fun make_markups () = global_markups := Symtab.foldl get_markup ([], ! global_parsers);
165 fun change_parsers f = (global_parsers := f (! global_parsers); make_markups ());
170 (* get current syntax *)
172 (*Note: the syntax for files is statically determined at the very
173 beginning; for interactive processing it may change dynamically.*)
175 fun get_lexicons () = ! global_lexicons;
176 fun get_parsers () = ! global_parsers;
177 fun get_parser () = apsome (#2 o #1) o curry Symtab.lookup (! global_parsers);
179 fun lookup_markup name = assoc (! global_markups, name);
180 fun is_markup kind name = (case lookup_markup name of Some k => k = kind | None => false);
185 fun add_keywords keywords = change_lexicons (apfst (fn lex =>
186 (Scan.extend_lexicon lex (map Symbol.explode keywords))));
188 fun add_parser (tab, Parser (name, (comment, kind, markup), int_only, parse)) =
189 (if is_none (Symtab.lookup (tab, name)) then ()
190 else warning ("Redefined outer syntax command " ^ quote name);
191 Symtab.update ((name, (((comment, kind), (int_only, parse)), markup)), tab));
193 fun add_parsers parsers =
194 (change_parsers (fn tab => foldl add_parser (tab, parsers));
195 change_lexicons (apsnd (fn lex => Scan.extend_lexicon lex
196 (map (fn Parser (name, _, _, _) => Symbol.explode name) parsers))));
203 fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ()));
205 fun dest_parsers () =
206 map (fn (name, (((cmt, kind), (int_only, _)), _)) => (name, cmt, kind, int_only))
207 (Symtab.dest (get_parsers ()));
209 fun help_outer_syntax () =
211 fun pretty_cmd (name, comment, _, _) =
212 Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
213 val (int_cmds, cmds) = partition #4 (dest_parsers ());
215 [Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())),
216 Pretty.big_list "proper commands:" (map pretty_cmd cmds),
217 Pretty.big_list "improper commands (interactive-only):" (map pretty_cmd int_cmds)]
220 val print_outer_syntax = Pretty.writeln o Pretty.chunks o help_outer_syntax;
223 Toplevel.keep (fn state =>
224 let val opt_thy = try Toplevel.theory_of state in
225 help_outer_syntax () @
226 Method.help_methods opt_thy @
227 Attrib.help_attributes opt_thy
228 |> Pretty.chunks |> Pretty.writeln
235 (* special keywords *)
237 val headerN = "header";
238 val theoryN = "theory";
240 val theory_keyword = P.$$$ theoryN;
241 val header_keyword = P.$$$ headerN;
242 val semicolon = P.$$$ ";";
250 Scan.unless semicolon (Scan.one (T.not_sync andf T.not_eof));
252 val recover = Scan.prompt "recover# " (Scan.repeat no_terminator);
256 fun source term do_recover cmd src =
258 |> Source.source T.stopper
259 (Scan.bulk (fn xs => P.!!! (command term (cmd ())) xs))
260 (if do_recover then Some recover else None)
261 |> Source.mapfilter I;
265 fun token_source (src, pos) =
267 |> Symbol.source false
268 |> T.source false (K (get_lexicons ())) pos;
270 fun filter_proper src =
272 |> Source.filter T.is_proper;
277 fun scan_header get_lex scan (src, pos) =
279 |> Symbol.source false
280 |> T.source false (fn () => (get_lex (), Scan.empty_lexicon)) pos
282 |> Source.source T.stopper (Scan.single scan) None
283 |> (fst o the o Source.get_single);
286 (* detect new/old header *)
290 val check_header_lexicon = Scan.make_lexicon [Symbol.explode headerN, Symbol.explode theoryN];
291 val check_header = Scan.option (header_keyword || theory_keyword);
295 fun is_old_theory src = is_none (scan_header (K check_header_lexicon) check_header src);
300 (* deps_thy --- inspect theory header *)
305 Scan.make_lexicon (map Symbol.explode ["(", ")", "+", ":", ";", "=", "files", headerN, theoryN]);
308 (P.$$$ "(" |-- P.!!! (P.name --| P.$$$ ")")) >> rpair false || P.name >> rpair true;
313 (P.name -- (P.$$$ "=" |-- P.enum1 "+" P.name) --
314 Scan.optional (P.$$$ "files" |-- P.!!! (Scan.repeat1 file_name)) [] --| P.$$$ ":")
315 >> (fn ((A, Bs), files) => (A, Bs, files));
318 header_keyword |-- (P.!!! (P.text -- Scan.option semicolon -- theory_keyword |-- theory_header))
319 || theory_keyword |-- P.!!! theory_header;
322 P.!!! (P.group "theory header"
323 (P.name -- (P.$$$ "=" |-- P.name -- Scan.repeat (P.$$$ "+" |-- P.name))))
324 >> (fn (A, (B, Bs)) => (A, B :: Bs, []: (string * bool) list));
326 fun deps_thy name ml path =
328 val src = Source.of_string (File.read path);
329 val pos = Path.position path;
330 val (name', parents, files) =
331 (*unfortunately, old-style headers dynamically depend on the current lexicon*)
332 if is_old_theory (src, pos) then
333 scan_header ThySyn.get_lexicon (Scan.error old_header) (src, pos)
334 else scan_header (K header_lexicon) (Scan.error new_header) (src, pos);
336 val ml_path = ThyLoad.ml_path name;
337 val ml_file = if ml andalso is_some (ThyLoad.check_file ml_path) then [ml_path] else [];
339 if name <> name' then
340 error ("Filename " ^ quote (Path.pack path) ^
341 " does not agree with theory name " ^ quote name')
342 else (parents, map (Path.unpack o #1) files @ ml_file)
348 (* present theory source *)
352 val is_improper = not o (T.is_proper orf T.is_begin_ignore orf T.is_end_ignore);
353 val improper = Scan.any is_improper;
355 val improper_keep_indent = Scan.repeat
356 (Scan.unless (Scan.one T.is_indent -- Scan.one T.is_proper) (Scan.one is_improper));
359 (improper -- semicolon) |-- improper_keep_indent ||
360 improper_keep_indent;
364 Scan.depend (fn d => Scan.one T.is_begin_ignore >> pair (d + 1)) ||
365 Scan.depend (fn 0 => Scan.fail | d => Scan.one T.is_end_ignore >> pair (d - 1)) ||
366 Scan.lift (Scan.one (OuterLex.not_eof andf (not o OuterLex.is_end_ignore)));
368 val opt_newline = Scan.option (Scan.one T.is_newline);
371 opt_newline -- Scan.one T.is_begin_ignore --
372 P.!!!! (Scan.pass 0 (Scan.repeat ignore) -- Scan.one T.is_end_ignore -- opt_newline);
374 fun markup_kind k = Scan.one (T.is_kind T.Command andf is_markup k o T.val_of);
376 val markup = markup_kind Markup >> T.val_of;
377 val markup_env = markup_kind MarkupEnv >> T.val_of;
378 val verbatim = markup_kind Verbatim;
381 ignore_stuff >> K None ||
382 (improper |-- markup -- P.!!!! (improper |-- P.text --| improper_end) >> Present.markup_token ||
383 improper |-- markup_env -- P.!!!! (improper |-- P.text --| improper_end) >> Present.markup_env_token ||
384 (P.$$$ "--" >> K "cmt") -- P.!!!! (improper |-- P.text) >> Present.markup_token ||
385 (improper -- verbatim) |-- P.!!!! (improper |-- P.text --| improper_end)
386 >> Present.verbatim_token ||
387 Scan.one T.not_eof >> Present.basic_token) >> Some;
391 (*note: lazy evaluation ahead*)
393 fun present_toks text pos () =
394 token_source (Source.of_list (Library.untabify text), pos)
395 |> Source.source T.stopper (Scan.bulk (Scan.error present_token)) None
396 |> Source.mapfilter I
399 fun present_text text () =
400 Source.exhaust (Symbol.source false (Source.of_list (Library.untabify text)));
405 (* load_thy --- read text (including header) *)
409 fun try_ml_file name time =
411 val path = ThyLoad.ml_path name;
412 val tr = Toplevel.imperative (fn () => ThyInfo.load_file time path);
413 val tr_name = if time then "time_use" else "use";
415 if is_none (ThyLoad.check_file path) then ()
416 else Toplevel.excursion_error [Toplevel.empty |> Toplevel.name tr_name |> tr]
419 fun parse_thy src_pos =
423 |> source false false (K (get_parser ()))
426 fun run_thy name path =
428 val text = explode (File.read path);
429 val src = Source.of_list text;
430 val pos = Path.position path;
432 Present.init_theory name;
433 Present.verbatim_source name (present_text text);
434 if is_old_theory (src, pos) then (ThySyn.load_thy name text;
435 Present.old_symbol_source name (present_text text)) (*note: text presented twice!*)
436 else (Toplevel.excursion_error (parse_thy (src, pos));
437 Present.token_source name (present_toks text pos))
442 fun load_thy name ml time path =
445 (writeln ("\n**** Starting theory " ^ quote name ^ " ****");
446 setmp Library.timing true (run_thy name) path;
447 writeln ("**** Finished theory " ^ quote name ^ " ****\n")))
448 else run_thy name path;
449 Context.context (ThyInfo.get_theory name);
450 if ml then try_ml_file name time else ());
455 (* interactive source of state transformers *)
457 fun isar term no_pos =
459 |> Symbol.source true
460 |> T.source true get_lexicons
461 (if no_pos then Position.none else Position.line_name 1 "stdin")
463 |> source term true get_parser;
467 (** the read-eval-print loop **)
471 fun gen_loop term no_pos =
472 (Context.reset_context ();
473 Toplevel.loop (isar term no_pos));
475 fun gen_main term no_pos =
476 (Toplevel.set_state Toplevel.toplevel;
477 writeln (Session.welcome ());
478 gen_loop term no_pos);
480 fun main () = gen_main false false;
481 fun loop () = gen_loop false false;
482 fun sync_main () = gen_main true true;
483 fun sync_loop () = gen_loop true true;
489 writeln ("This is Isabelle's underlying ML system (" ^ ml_system ^ ");\n\
490 \invoke 'loop();' to enter the Isar loop.");
493 (*final declarations of this structure!*)
494 val command = parser false None;
495 val markup_command = parser false (Some Markup);
496 val markup_env_command = parser false (Some MarkupEnv);
497 val verbatim_command = parser false (Some Verbatim);
498 val improper_command = parser true None;
503 (*setup theory syntax dependent operations*)
504 ThyLoad.deps_thy_fn := OuterSyntax.deps_thy;
505 ThyLoad.load_thy_fn := OuterSyntax.load_thy;
506 structure ThyLoad: THY_LOAD = ThyLoad;
508 structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax;
509 open BasicOuterSyntax;