5829
|
1 |
(* Title: Pure/Isar/outer_syntax.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Markus Wenzel, TU Muenchen
|
|
4 |
|
|
5 |
The global Isabelle/Isar outer syntax.
|
|
6 |
*)
|
|
7 |
|
|
8 |
signature BASIC_OUTER_SYNTAX =
|
|
9 |
sig
|
|
10 |
val main: unit -> unit
|
5883
|
11 |
val loop: unit -> unit
|
6860
|
12 |
val sync_main: unit -> unit
|
|
13 |
val sync_loop: unit -> unit
|
5829
|
14 |
val help: unit -> unit
|
|
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
|
|
25 |
val thy_end: string
|
|
26 |
val thy_heading: string
|
|
27 |
val thy_decl: string
|
|
28 |
val thy_goal: string
|
|
29 |
val qed: string
|
6733
|
30 |
val qed_block: string
|
6722
|
31 |
val prf_goal: string
|
|
32 |
val prf_block: string
|
|
33 |
val prf_chain: string
|
|
34 |
val prf_decl: string
|
6868
|
35 |
val prf_asm: string
|
6722
|
36 |
val prf_script: string
|
|
37 |
val kinds: string list
|
|
38 |
end
|
5829
|
39 |
type token
|
|
40 |
type parser
|
6722
|
41 |
val command: string -> string -> string ->
|
6373
|
42 |
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
|
6722
|
43 |
val improper_command: string -> string -> string ->
|
6373
|
44 |
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser
|
7026
|
45 |
val dest_keywords: unit -> string list
|
|
46 |
val dest_parsers: unit -> (string * string * string * bool) list
|
5883
|
47 |
val print_outer_syntax: unit -> unit
|
5829
|
48 |
val add_keywords: string list -> unit
|
|
49 |
val add_parsers: parser list -> unit
|
6247
|
50 |
val theory_header: token list -> (string * string list * (string * bool) list) * token list
|
6199
|
51 |
val deps_thy: string -> bool -> Path.T -> string list * Path.T list
|
|
52 |
val load_thy: string -> bool -> bool -> Path.T -> unit
|
6860
|
53 |
val isar: bool -> Toplevel.isar
|
5829
|
54 |
end;
|
|
55 |
|
|
56 |
structure OuterSyntax: OUTER_SYNTAX =
|
|
57 |
struct
|
|
58 |
|
6860
|
59 |
structure P = OuterParse;
|
|
60 |
|
5829
|
61 |
|
|
62 |
(** outer syntax **)
|
|
63 |
|
6722
|
64 |
(* command keyword classification *)
|
|
65 |
|
|
66 |
structure Keyword =
|
|
67 |
struct
|
|
68 |
val control = "control";
|
|
69 |
val diag = "diag";
|
|
70 |
val thy_begin = "theory-begin";
|
|
71 |
val thy_end = "theory-end";
|
|
72 |
val thy_heading = "theory-heading";
|
|
73 |
val thy_decl = "theory-decl";
|
|
74 |
val thy_goal = "theory-goal";
|
|
75 |
val qed = "qed";
|
6733
|
76 |
val qed_block = "qed-block";
|
6722
|
77 |
val prf_goal = "proof-goal";
|
|
78 |
val prf_block = "proof-block";
|
|
79 |
val prf_chain = "proof-chain";
|
|
80 |
val prf_decl = "proof-decl";
|
6868
|
81 |
val prf_asm = "proof-asm";
|
6722
|
82 |
val prf_script = "proof-script";
|
|
83 |
|
|
84 |
val kinds = [control, diag, thy_begin, thy_end, thy_heading, thy_decl, thy_goal, qed,
|
6868
|
85 |
qed_block, prf_goal, prf_block, prf_chain, prf_decl, prf_asm, prf_script];
|
6722
|
86 |
end;
|
|
87 |
|
|
88 |
|
5829
|
89 |
(* parsers *)
|
|
90 |
|
|
91 |
type token = OuterLex.token;
|
|
92 |
type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list;
|
|
93 |
|
|
94 |
datatype parser =
|
6722
|
95 |
Parser of string * (string * string) * bool * parser_fn;
|
5829
|
96 |
|
6722
|
97 |
fun parser int_only name comment kind parse = Parser (name, (comment, kind), int_only, parse);
|
5829
|
98 |
|
|
99 |
|
|
100 |
(* parse command *)
|
|
101 |
|
6860
|
102 |
local
|
6199
|
103 |
|
5829
|
104 |
fun command_body cmd (name, _) =
|
7026
|
105 |
(case cmd name of
|
|
106 |
Some (int_only, parse) => P.!!! (Scan.prompt (name ^ "# ") (parse >> pair int_only))
|
|
107 |
| None => sys_error ("no parser for outer syntax command " ^ quote name));
|
6860
|
108 |
|
|
109 |
fun terminator false = Scan.succeed ()
|
|
110 |
| terminator true = P.group "terminator" (Scan.option P.sync -- P.$$$ ";" >> K ());
|
|
111 |
|
|
112 |
in
|
5829
|
113 |
|
6860
|
114 |
fun command term cmd =
|
|
115 |
P.$$$ ";" >> K None ||
|
|
116 |
P.sync >> K None ||
|
7026
|
117 |
(P.position P.command :-- command_body cmd) --| terminator term
|
6860
|
118 |
>> (fn ((name, pos), (int_only, f)) =>
|
|
119 |
Some (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
|
|
120 |
Toplevel.interactive int_only |> f));
|
5829
|
121 |
|
6199
|
122 |
end;
|
|
123 |
|
5829
|
124 |
|
|
125 |
|
|
126 |
(** global syntax state **)
|
|
127 |
|
7026
|
128 |
local
|
|
129 |
|
|
130 |
val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon);
|
6722
|
131 |
val global_parsers = ref (Symtab.empty: ((string * string) * (bool * parser_fn)) Symtab.table);
|
5952
|
132 |
|
7026
|
133 |
fun change_lexicons f =
|
|
134 |
let val lexs = f (! global_lexicons) in
|
|
135 |
(case (op inter_string) (pairself Scan.dest_lexicon lexs) of
|
|
136 |
[] => global_lexicons := lexs
|
|
137 |
| bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads))
|
|
138 |
end;
|
5829
|
139 |
|
7026
|
140 |
fun change_parsers f = global_parsers := f (! global_parsers);
|
6722
|
141 |
|
7026
|
142 |
in
|
|
143 |
|
|
144 |
(* get current lexers / parsers *)
|
|
145 |
|
|
146 |
(*Note: the syntax for files is statically determined at the very
|
|
147 |
beginning; for interactive processing it may change dynamically.*)
|
|
148 |
|
|
149 |
fun get_lexicons () = ! global_lexicons;
|
|
150 |
fun get_parsers () = ! global_parsers;
|
|
151 |
fun get_parser () = apsome snd o curry Symtab.lookup (! global_parsers);
|
5829
|
152 |
|
|
153 |
|
|
154 |
(* augment syntax *)
|
|
155 |
|
7026
|
156 |
fun add_keywords keywords = change_lexicons (apfst (fn lex =>
|
|
157 |
(Scan.extend_lexicon lex (map Symbol.explode keywords))));
|
5829
|
158 |
|
|
159 |
fun add_parser (tab, Parser (name, comment, int_only, parse)) =
|
|
160 |
(if is_none (Symtab.lookup (tab, name)) then ()
|
|
161 |
else warning ("Redefined outer syntax command " ^ quote name);
|
|
162 |
Symtab.update ((name, (comment, (int_only, parse))), tab));
|
|
163 |
|
|
164 |
fun add_parsers parsers =
|
7026
|
165 |
(change_parsers (fn tab => foldl add_parser (tab, parsers));
|
|
166 |
change_lexicons (apsnd (fn lex => Scan.extend_lexicon lex
|
|
167 |
(map (fn Parser (name, _, _, _) => Symbol.explode name) parsers))));
|
|
168 |
|
|
169 |
end;
|
5829
|
170 |
|
|
171 |
|
7026
|
172 |
(* print syntax *)
|
|
173 |
|
|
174 |
fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ()));
|
|
175 |
|
|
176 |
fun dest_parsers () =
|
|
177 |
map (fn (name, ((cmt, kind), (int_only, _))) => (name, cmt, kind, int_only))
|
|
178 |
(Symtab.dest (get_parsers ()));
|
5829
|
179 |
|
7026
|
180 |
fun print_outer_syntax () =
|
|
181 |
let
|
|
182 |
fun pretty_cmd (name, comment, _, _) =
|
|
183 |
Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
|
|
184 |
val (int_cmds, cmds) = partition #4 (dest_parsers ());
|
|
185 |
in
|
|
186 |
Pretty.writeln (Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())));
|
|
187 |
Pretty.writeln (Pretty.big_list "proper commands:" (map pretty_cmd cmds));
|
|
188 |
Pretty.writeln (Pretty.big_list "improper commands (interactive-only):"
|
|
189 |
(map pretty_cmd int_cmds))
|
|
190 |
end;
|
5829
|
191 |
|
|
192 |
|
|
193 |
|
|
194 |
(** read theory **)
|
|
195 |
|
6247
|
196 |
(* theory keyword *)
|
|
197 |
|
|
198 |
val theoryN = "theory";
|
|
199 |
val theory_keyword = OuterParse.$$$ theoryN;
|
|
200 |
|
|
201 |
|
5829
|
202 |
(* source *)
|
|
203 |
|
6860
|
204 |
local
|
5829
|
205 |
|
6860
|
206 |
val no_terminator =
|
|
207 |
Scan.unless (P.$$$ ";") (Scan.one (OuterLex.not_sync andf OuterLex.not_eof));
|
|
208 |
|
|
209 |
val recover = Scan.prompt "recover# " (Scan.repeat no_terminator);
|
|
210 |
|
|
211 |
in
|
5829
|
212 |
|
6860
|
213 |
fun source term do_recover cmd src =
|
5829
|
214 |
src
|
6860
|
215 |
|> Source.source OuterLex.stopper
|
|
216 |
(Scan.bulk (fn xs => OuterParse.!!! (command term (cmd ())) xs))
|
|
217 |
(if do_recover then Some recover else None)
|
5829
|
218 |
|> Source.mapfilter I;
|
|
219 |
|
6860
|
220 |
end;
|
|
221 |
|
5829
|
222 |
|
|
223 |
(* detect header *)
|
|
224 |
|
7026
|
225 |
fun scan_header get_lex scan (src, pos) =
|
5829
|
226 |
src
|
|
227 |
|> Symbol.source false
|
7026
|
228 |
|> OuterLex.source false (fn () => (get_lex (), Scan.empty_lexicon)) pos
|
6199
|
229 |
|> Source.source OuterLex.stopper (Scan.single scan) None
|
5829
|
230 |
|> (fst o the o Source.get_single);
|
|
231 |
|
6247
|
232 |
val check_header_lexicon = Scan.make_lexicon [Symbol.explode theoryN];
|
5829
|
233 |
|
6199
|
234 |
fun is_old_theory src =
|
6247
|
235 |
is_none (scan_header (K check_header_lexicon) (Scan.option theory_keyword) src);
|
6199
|
236 |
|
|
237 |
fun warn_theory_style path is_old =
|
|
238 |
let
|
|
239 |
val style = if is_old then "old" else "new";
|
|
240 |
val _ = warning ("Assuming " ^ style ^ "-style theory format for " ^ quote (Path.pack path));
|
|
241 |
in is_old end;
|
|
242 |
|
|
243 |
|
|
244 |
(* deps_thy --- inspect theory header *)
|
|
245 |
|
6247
|
246 |
val header_lexicon =
|
|
247 |
Scan.make_lexicon (map Symbol.explode ["(", ")", "+", ":", "=", "files", theoryN]);
|
6199
|
248 |
|
6860
|
249 |
local
|
6199
|
250 |
|
7026
|
251 |
val file_name =
|
|
252 |
(P.$$$ "(" |-- P.!!! (P.name --| P.$$$ ")")) >> rpair false || P.name >> rpair true;
|
6247
|
253 |
|
|
254 |
val theory_head =
|
6860
|
255 |
(P.name -- (P.$$$ "=" |-- P.enum1 "+" P.name) --
|
|
256 |
Scan.optional (P.$$$ "files" |-- P.!!! (Scan.repeat1 file_name)) [])
|
6247
|
257 |
>> (fn ((A, Bs), files) => (A, Bs, files));
|
|
258 |
|
6860
|
259 |
in
|
|
260 |
|
|
261 |
val theory_header = theory_head --| (Scan.ahead P.eof || P.$$$ ":");
|
|
262 |
val only_header = theory_keyword |-- theory_head --| Scan.ahead P.eof;
|
|
263 |
val new_header = theory_keyword |-- P.!!! theory_header;
|
6199
|
264 |
|
|
265 |
val old_header =
|
6860
|
266 |
P.name -- (P.$$$ "=" |-- P.name -- Scan.repeat (P.$$$ "+" |-- P.name))
|
6247
|
267 |
>> (fn (A, (B, Bs)) => (A, B :: Bs, []: (string * bool) list));
|
6199
|
268 |
|
|
269 |
end;
|
5829
|
270 |
|
6199
|
271 |
fun deps_thy name ml path =
|
|
272 |
let
|
6641
|
273 |
val src = File.source path;
|
6199
|
274 |
val is_old = warn_theory_style path (is_old_theory src);
|
6247
|
275 |
val (name', parents, files) =
|
6199
|
276 |
(*Note: old style headers dynamically depend on the current lexicon :-( *)
|
|
277 |
if is_old then scan_header ThySyn.get_lexicon (Scan.error old_header) src
|
6247
|
278 |
else scan_header (K header_lexicon) (Scan.error new_header) src;
|
6199
|
279 |
|
|
280 |
val ml_path = ThyLoad.ml_path name;
|
|
281 |
val ml_file = if not ml orelse is_none (ThyLoad.check_file ml_path) then [] else [ml_path];
|
|
282 |
in
|
|
283 |
if name <> name' then
|
|
284 |
error ("Filename " ^ quote (Path.pack path) ^ " does not match theory name " ^ quote name)
|
6247
|
285 |
else (parents, map (Path.unpack o #1) files @ ml_file)
|
6199
|
286 |
end;
|
|
287 |
|
|
288 |
|
|
289 |
(* load_thy --- read text (including header) *)
|
|
290 |
|
6247
|
291 |
fun try_ml_file name ml time =
|
6199
|
292 |
let
|
|
293 |
val path = ThyLoad.ml_path name;
|
6247
|
294 |
val tr = Toplevel.imperative (fn () => ThyInfo.load_file time path);
|
|
295 |
val tr_name = if time then "time_use" else "use";
|
6199
|
296 |
in
|
|
297 |
if not ml orelse is_none (ThyLoad.check_file path) then ()
|
6247
|
298 |
else Toplevel.excursion [Toplevel.empty |> Toplevel.name tr_name |> tr]
|
6199
|
299 |
end;
|
|
300 |
|
|
301 |
fun parse_thy (src, pos) =
|
6247
|
302 |
let
|
|
303 |
val lex_src =
|
|
304 |
src
|
|
305 |
|> Symbol.source false
|
7026
|
306 |
|> OuterLex.source false (K (get_lexicons ())) pos;
|
6247
|
307 |
val only_head =
|
|
308 |
lex_src
|
|
309 |
|> Source.source OuterLex.stopper (Scan.single (Scan.option only_header)) None
|
|
310 |
|> (fst o the o Source.get_single);
|
|
311 |
in
|
|
312 |
(case only_head of
|
|
313 |
None =>
|
|
314 |
lex_src
|
6860
|
315 |
|> source false false (K (get_parser ()))
|
6247
|
316 |
|> Source.exhaust
|
|
317 |
| Some spec =>
|
|
318 |
[Toplevel.empty |> Toplevel.name theoryN |> IsarThy.theory spec,
|
|
319 |
Toplevel.empty |> Toplevel.name "end" |> Toplevel.exit])
|
|
320 |
end;
|
5829
|
321 |
|
6247
|
322 |
fun run_thy name path =
|
6641
|
323 |
let val (src, pos) = File.source path in
|
6332
|
324 |
Present.theory_source name src;
|
6247
|
325 |
if is_old_theory (src, pos) then ThySyn.load_thy name (Source.exhaust src)
|
|
326 |
else (Toplevel.excursion (parse_thy (src, pos))
|
|
327 |
handle exn => error (Toplevel.exn_message exn))
|
|
328 |
end;
|
6199
|
329 |
|
|
330 |
fun load_thy name ml time path =
|
6247
|
331 |
(if time then
|
|
332 |
timeit (fn () =>
|
|
333 |
(writeln ("\n**** Starting theory " ^ quote name ^ " ****");
|
|
334 |
setmp Goals.proof_timing true (run_thy name) path;
|
|
335 |
writeln ("**** Finished theory " ^ quote name ^ " ****\n")))
|
|
336 |
else run_thy name path;
|
|
337 |
Context.context (ThyInfo.get_theory name);
|
|
338 |
try_ml_file name ml time);
|
5829
|
339 |
|
|
340 |
|
|
341 |
(* interactive source of state transformers *)
|
|
342 |
|
6860
|
343 |
fun isar term =
|
5829
|
344 |
Source.tty
|
|
345 |
|> Symbol.source true
|
7026
|
346 |
|> OuterLex.source true get_lexicons (Position.line_name 1 "stdin")
|
6860
|
347 |
|> source term true get_parser;
|
5829
|
348 |
|
|
349 |
|
|
350 |
|
|
351 |
(** the read-eval-print loop **)
|
|
352 |
|
5923
|
353 |
(* main loop *)
|
|
354 |
|
6860
|
355 |
fun gen_loop term = (Context.reset_context (); Toplevel.loop (isar term));
|
5829
|
356 |
|
6860
|
357 |
fun gen_main term =
|
5829
|
358 |
(Toplevel.set_state Toplevel.toplevel;
|
|
359 |
ml_prompts "ML> " "ML# ";
|
6199
|
360 |
writeln (Session.welcome ());
|
6860
|
361 |
gen_loop term);
|
|
362 |
|
|
363 |
fun main () = gen_main false;
|
|
364 |
fun loop () = gen_loop false;
|
|
365 |
fun sync_main () = gen_main true;
|
|
366 |
fun sync_loop () = gen_loop true;
|
5829
|
367 |
|
|
368 |
|
|
369 |
(* help *)
|
|
370 |
|
|
371 |
fun help () =
|
|
372 |
writeln ("This is Isabelle's underlying ML system (" ^ ml_system ^ ");\n\
|
5883
|
373 |
\invoke 'loop();' to enter the Isar loop.");
|
5829
|
374 |
|
|
375 |
|
6373
|
376 |
(*final declarations of this structure!*)
|
6685
|
377 |
val command = parser false;
|
|
378 |
val improper_command = parser true;
|
|
379 |
|
6373
|
380 |
|
5829
|
381 |
end;
|
|
382 |
|
6199
|
383 |
(*setup theory syntax dependent operations*)
|
|
384 |
ThyLoad.deps_thy_fn := OuterSyntax.deps_thy;
|
|
385 |
ThyLoad.load_thy_fn := OuterSyntax.load_thy;
|
|
386 |
structure ThyLoad: THY_LOAD = ThyLoad;
|
|
387 |
|
5829
|
388 |
structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax;
|
|
389 |
open BasicOuterSyntax;
|