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
17 signature OUTER_SYNTAX =
19 include BASIC_OUTER_SYNTAX
25 val thy_switch: string
27 val thy_heading: string
29 val thy_script: string
33 val qed_global: string
34 val prf_heading: string
42 val prf_asm_goal: string
43 val prf_script: string
44 val kinds: string list
48 val command: string -> string -> string ->
49 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
50 val markup_command: IsarOutput.markup -> string -> string -> string ->
51 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
52 val improper_command: string -> string -> string ->
53 (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
54 val is_keyword: string -> bool
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_commands: Toplevel.transition -> Toplevel.transition
59 val add_keywords: string list -> unit
60 val add_parsers: parser list -> unit
61 val check_text: string * Position.T -> bool -> Toplevel.state -> unit
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 -> unit Toplevel.isar
65 val isar_readstring: Position.T -> string -> (string list) Toplevel.isar
66 val read: string -> (string * OuterLex.token list * Toplevel.transition) list
69 structure OuterSyntax: OUTER_SYNTAX =
72 structure T = OuterLex;
73 structure P = OuterParse;
78 (* command keyword classification *)
82 val control = "control";
84 val thy_begin = "theory-begin";
85 val thy_switch = "theory-switch";
86 val thy_end = "theory-end";
87 val thy_heading = "theory-heading";
88 val thy_decl = "theory-decl";
89 val thy_script = "theory-script";
90 val thy_goal = "theory-goal";
92 val qed_block = "qed-block";
93 val qed_global = "qed-global";
94 val prf_heading = "proof-heading";
95 val prf_goal = "proof-goal";
96 val prf_block = "proof-block";
97 val prf_open = "proof-open";
98 val prf_close = "proof-close";
99 val prf_chain = "proof-chain";
100 val prf_decl = "proof-decl";
101 val prf_asm = "proof-asm";
102 val prf_asm_goal = "proof-asm-goal";
103 val prf_script = "proof-script";
105 val kinds = [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script,
106 thy_goal, qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open, prf_close,
107 prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script];
113 type token = T.token;
114 type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list;
117 Parser of string * (string * string * IsarOutput.markup option) * bool * parser_fn;
119 fun parser int_only markup name comment kind parse =
120 Parser (name, (comment, kind, markup), int_only, parse);
127 fun terminate false = Scan.succeed ()
128 | terminate true = P.group "end of input" (Scan.option P.sync -- P.semicolon >> K ());
130 fun trace false parse = parse
131 | trace true parse = Scan.trace parse >> (fn (f, toks) => f o Toplevel.source toks);
133 fun body cmd trc (name, _) =
135 Some (int_only, parse) =>
136 P.!!! (Scan.prompt (name ^ "# ") (trace trc parse >> pair int_only))
137 | None => sys_error ("no parser for outer syntax command " ^ quote name));
141 fun command do_terminate do_trace cmd =
142 P.semicolon >> K None ||
144 (P.position P.command :-- body cmd do_trace) --| terminate do_terminate
145 >> (fn ((name, pos), (int_only, f)) =>
146 Some (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
147 Toplevel.interactive int_only |> f));
153 (** global outer syntax **)
157 val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon);
159 ref (Symtab.empty: (((string * string) * (bool * parser_fn)) * IsarOutput.markup option)
161 val global_markups = ref ([]: (string * IsarOutput.markup) list);
163 fun change_lexicons f =
164 let val lexs = f (! global_lexicons) in
165 (case (op inter_string) (pairself Scan.dest_lexicon lexs) of
166 [] => global_lexicons := lexs
167 | bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads))
170 fun get_markup (ms, (name, (_, Some m))) = (name, m) :: ms
171 | get_markup (ms, _) = ms;
173 fun make_markups () = global_markups := Symtab.foldl get_markup ([], ! global_parsers);
174 fun change_parsers f = (Library.change global_parsers f; make_markups ());
179 (* access current syntax *)
181 (*Note: the syntax for files is statically determined at the very
182 beginning; for interactive processing it may change dynamically.*)
184 fun get_lexicons () = ! global_lexicons;
185 fun get_parsers () = ! global_parsers;
186 fun get_parser () = apsome (#2 o #1) o curry Symtab.lookup (! global_parsers);
188 fun is_markup kind name =
189 (case assoc (! global_markups, name) of Some k => k = kind | None => false);
190 fun markup kind = Scan.one (T.is_kind T.Command andf is_markup kind o T.val_of);
195 fun add_keywords keywords = change_lexicons (apfst (fn lex =>
196 (Scan.extend_lexicon lex (map Symbol.explode keywords))));
198 fun add_parser (tab, Parser (name, (comment, kind, markup), int_only, parse)) =
199 (if is_none (Symtab.lookup (tab, name)) then ()
200 else warning ("Redefined outer syntax command " ^ quote name);
201 Symtab.update ((name, (((comment, kind), (int_only, parse)), markup)), tab));
203 fun add_parsers parsers =
204 (change_parsers (fn tab => foldl add_parser (tab, parsers));
205 change_lexicons (apsnd (fn lex => Scan.extend_lexicon lex
206 (map (fn Parser (name, _, _, _) => Symbol.explode name) parsers))));
213 fun is_keyword s = Scan.is_literal (#1 (get_lexicons ())) (Symbol.explode s);
214 fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ()));
216 fun dest_parsers () =
217 map (fn (name, (((cmt, kind), (int_only, _)), _)) => (name, cmt, kind, int_only))
218 (Symtab.dest (get_parsers ()));
220 fun print_outer_syntax () =
222 fun pretty_cmd (name, comment, _, _) =
223 Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
224 val (int_cmds, cmds) = partition #4 (dest_parsers ());
226 [Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())),
227 Pretty.big_list "proper commands:" (map pretty_cmd cmds),
228 Pretty.big_list "improper commands (interactive-only):" (map pretty_cmd int_cmds)]
229 |> Pretty.chunks |> Pretty.writeln
232 val print_commands = Toplevel.imperative print_outer_syntax;
236 (** toplevel parsing **)
240 fun toplevel_source term trc do_recover cmd src =
243 Scan.unless P.semicolon (Scan.one (T.not_sync andf T.not_eof));
244 fun recover x = (Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [None]) x;
248 |> Source.source T.stopper
249 (Scan.bulk (P.$$$ "--" -- P.!!! P.text >> K None || P.not_eof >> Some))
250 (if do_recover then Some recover else None)
251 |> Source.mapfilter I
252 |> Source.source T.stopper (Scan.bulk (fn xs => P.!!! (command term trc (cmd ())) xs))
253 (if do_recover then Some recover else None)
254 |> Source.mapfilter I
258 (* interactive source of toplevel transformers *)
260 fun isar term no_pos =
262 |> Symbol.source true
263 |> T.source true get_lexicons
264 (if no_pos then Position.none else Position.line_name 1 "stdin")
265 |> toplevel_source term false true get_parser;
268 (* string source of transformers with trace (for Proof General) *)
270 fun isar_readstring pos str =
272 |> Symbol.source false
273 |> T.source false get_lexicons pos
274 |> toplevel_source false true true get_parser;
277 (* read text -- with trace of source *)
281 |> Symbol.source false
282 |> T.source false get_lexicons Position.none
283 |> toplevel_source false true true get_parser
285 |> map (fn tr => (Toplevel.name_of tr, the (Toplevel.source_of tr), tr));
293 fun check_text s true state = (IsarOutput.eval_antiquote (#1 (get_lexicons ())) state s; ())
294 | check_text _ false _ = ();
299 fun deps_thy name ml path =
301 val src = Source.of_string (File.read path);
302 val pos = Path.position path;
303 val (name', parents, files) = ThyHeader.scan (src, pos);
304 val ml_path = ThyLoad.ml_path name;
305 val ml_file = if ml andalso is_some (ThyLoad.check_file ml_path) then [ml_path] else [];
307 if name <> name' then
308 error ("Filename " ^ quote (Path.pack path) ^
309 " does not agree with theory name " ^ quote name')
310 else (parents, map (Path.unpack o #1) files @ ml_file)
318 fun try_ml_file name time =
320 val path = ThyLoad.ml_path name;
321 val tr = Toplevel.imperative (fn () => ThyInfo.load_file time path);
322 val tr_name = if time then "time_use" else "use";
324 if is_none (ThyLoad.check_file path) then ()
325 else Toplevel.excursion [Toplevel.empty |> Toplevel.name tr_name |> tr]
330 |> toplevel_source false false false (K (get_parser ()))
333 fun run_thy name path =
335 val pos = Path.position path;
336 val text = Library.untabify (explode (File.read path));
337 val text_src = Source.of_list text;
338 fun present_text () = Source.exhaust (Symbol.source false text_src);
340 Present.init_theory name;
341 Present.verbatim_source name present_text;
342 if ThyHeader.is_old (text_src, pos) then (ThySyn.load_thy name text;
343 Present.old_symbol_source name present_text) (*note: text presented twice*)
346 val tok_src = text_src
347 |> Symbol.source false
348 |> T.source false (K (get_lexicons ())) pos
350 val out = Toplevel.excursion_result
351 (IsarOutput.parse_thy markup (#1 (get_lexicons ()))
352 (parse_thy tok_src) tok_src);
353 in Present.theory_output name (Buffer.content out) end
358 fun load_thy name ml time path =
361 (writeln ("\n**** Starting theory " ^ quote name ^ " ****");
363 writeln ("**** Finished theory " ^ quote name ^ " ****\n")))
364 else run_thy name path;
365 Context.context (ThyInfo.get_theory name);
366 if ml then try_ml_file name time else ());
372 (** the read-eval-print loop **)
376 fun gen_loop term no_pos =
377 (Context.reset_context ();
378 Toplevel.loop (isar term no_pos));
380 fun gen_main term no_pos =
381 (Toplevel.set_state Toplevel.toplevel;
382 writeln (Session.welcome ());
383 gen_loop term no_pos);
385 fun main () = gen_main false false;
386 fun loop () = gen_loop false false;
387 fun sync_main () = gen_main true true;
388 fun sync_loop () = gen_loop true true;
394 writeln ("This is Isabelle's underlying ML system (" ^ ml_system ^ ");\n\
395 \invoke 'Isar.loop();' to get back to the Isar read-eval-print loop.");
398 (*final declarations of this structure!*)
399 val command = parser false None;
400 val markup_command = parser false o Some;
401 val improper_command = parser true None;
406 (*setup theory syntax dependent operations*)
407 ThyLoad.deps_thy_fn := OuterSyntax.deps_thy;
408 ThyLoad.load_thy_fn := OuterSyntax.load_thy;
409 structure ThyLoad: THY_LOAD = ThyLoad;
411 structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax;
412 open BasicOuterSyntax;