| author | paulson | 
| Mon, 22 May 2000 13:20:47 +0200 | |
| changeset 8919 | d00b01ed8539 | 
| parent 8807 | 0046be1769f9 | 
| child 8999 | ad8260dc6e4a | 
| permissions | -rw-r--r-- | 
| 5829 | 1  | 
(* Title: Pure/Isar/outer_syntax.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Markus Wenzel, TU Muenchen  | 
|
| 8807 | 4  | 
License: GPL (GNU GENERAL PUBLIC LICENSE)  | 
| 5829 | 5  | 
|
6  | 
The global Isabelle/Isar outer syntax.  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
signature BASIC_OUTER_SYNTAX =  | 
|
10  | 
sig  | 
|
11  | 
val main: unit -> unit  | 
|
| 5883 | 12  | 
val loop: unit -> unit  | 
| 6860 | 13  | 
val sync_main: unit -> unit  | 
14  | 
val sync_loop: unit -> unit  | 
|
| 5829 | 15  | 
val help: unit -> unit  | 
16  | 
end;  | 
|
17  | 
||
18  | 
signature OUTER_SYNTAX =  | 
|
19  | 
sig  | 
|
20  | 
include BASIC_OUTER_SYNTAX  | 
|
| 6722 | 21  | 
structure Keyword:  | 
22  | 
sig  | 
|
23  | 
val control: string  | 
|
24  | 
val diag: string  | 
|
25  | 
val thy_begin: string  | 
|
| 7104 | 26  | 
val thy_switch: string  | 
| 6722 | 27  | 
val thy_end: string  | 
28  | 
val thy_heading: string  | 
|
29  | 
val thy_decl: string  | 
|
30  | 
val thy_goal: string  | 
|
31  | 
val qed: string  | 
|
| 6733 | 32  | 
val qed_block: string  | 
| 8209 | 33  | 
val qed_global: string  | 
| 6722 | 34  | 
val prf_goal: string  | 
35  | 
val prf_block: string  | 
|
36  | 
val prf_chain: string  | 
|
37  | 
val prf_decl: string  | 
|
| 6868 | 38  | 
val prf_asm: string  | 
| 7676 | 39  | 
val prf_asm_goal: string  | 
| 6722 | 40  | 
val prf_script: string  | 
41  | 
val kinds: string list  | 
|
42  | 
end  | 
|
| 5829 | 43  | 
type token  | 
44  | 
type parser  | 
|
| 6722 | 45  | 
val command: string -> string -> string ->  | 
| 6373 | 46  | 
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser  | 
| 7750 | 47  | 
val markup_command: string -> string -> string ->  | 
48  | 
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser  | 
|
| 8660 | 49  | 
val markup_env_command: string -> string -> string ->  | 
50  | 
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser  | 
|
| 7789 | 51  | 
val verbatim_command: string -> string -> string ->  | 
52  | 
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser  | 
|
| 6722 | 53  | 
val improper_command: string -> string -> string ->  | 
| 6373 | 54  | 
(token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser  | 
| 7026 | 55  | 
val dest_keywords: unit -> string list  | 
56  | 
val dest_parsers: unit -> (string * string * string * bool) list  | 
|
| 5883 | 57  | 
val print_outer_syntax: unit -> unit  | 
| 7367 | 58  | 
val print_help: Toplevel.transition -> Toplevel.transition  | 
| 5829 | 59  | 
val add_keywords: string list -> unit  | 
60  | 
val add_parsers: parser list -> unit  | 
|
| 6247 | 61  | 
val theory_header: token list -> (string * string list * (string * bool) list) * token list  | 
| 7940 | 62  | 
val deps_thy: string -> bool -> Path.T -> string list * Path.T list  | 
| 6199 | 63  | 
val load_thy: string -> bool -> bool -> Path.T -> unit  | 
| 7333 | 64  | 
val isar: bool -> bool -> Toplevel.isar  | 
| 5829 | 65  | 
end;  | 
66  | 
||
67  | 
structure OuterSyntax: OUTER_SYNTAX =  | 
|
68  | 
struct  | 
|
69  | 
||
| 7750 | 70  | 
structure T = OuterLex;  | 
| 6860 | 71  | 
structure P = OuterParse;  | 
72  | 
||
| 5829 | 73  | 
|
74  | 
(** outer syntax **)  | 
|
75  | 
||
| 6722 | 76  | 
(* command keyword classification *)  | 
77  | 
||
78  | 
structure Keyword =  | 
|
79  | 
struct  | 
|
80  | 
val control = "control";  | 
|
81  | 
val diag = "diag";  | 
|
82  | 
val thy_begin = "theory-begin";  | 
|
| 7104 | 83  | 
val thy_switch = "theory-switch";  | 
| 6722 | 84  | 
val thy_end = "theory-end";  | 
85  | 
val thy_heading = "theory-heading";  | 
|
86  | 
val thy_decl = "theory-decl";  | 
|
87  | 
val thy_goal = "theory-goal";  | 
|
88  | 
val qed = "qed";  | 
|
| 6733 | 89  | 
val qed_block = "qed-block";  | 
| 8209 | 90  | 
val qed_global = "qed-global";  | 
| 6722 | 91  | 
val prf_goal = "proof-goal";  | 
92  | 
val prf_block = "proof-block";  | 
|
93  | 
val prf_chain = "proof-chain";  | 
|
94  | 
val prf_decl = "proof-decl";  | 
|
| 6868 | 95  | 
val prf_asm = "proof-asm";  | 
| 7676 | 96  | 
val prf_asm_goal = "proof-asm-goal";  | 
| 6722 | 97  | 
val prf_script = "proof-script";  | 
98  | 
||
| 7104 | 99  | 
val kinds = [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_goal,  | 
| 8209 | 100  | 
qed, qed_block, qed_global, prf_goal, prf_block, prf_chain, prf_decl, prf_asm, prf_asm_goal,  | 
101  | 
prf_script];  | 
|
| 6722 | 102  | 
end;  | 
103  | 
||
104  | 
||
| 5829 | 105  | 
(* parsers *)  | 
106  | 
||
| 7750 | 107  | 
type token = T.token;  | 
| 5829 | 108  | 
type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list;  | 
109  | 
||
| 8660 | 110  | 
datatype markup_kind = Markup | MarkupEnv | Verbatim;  | 
111  | 
||
| 5829 | 112  | 
datatype parser =  | 
| 8660 | 113  | 
Parser of string * (string * string * markup_kind option) * bool * parser_fn;  | 
| 5829 | 114  | 
|
| 7750 | 115  | 
fun parser int_only markup name comment kind parse =  | 
116  | 
Parser (name, (comment, kind, markup), int_only, parse);  | 
|
| 5829 | 117  | 
|
118  | 
||
119  | 
(* parse command *)  | 
|
120  | 
||
| 6860 | 121  | 
local  | 
| 6199 | 122  | 
|
| 5829 | 123  | 
fun command_body cmd (name, _) =  | 
| 7026 | 124  | 
(case cmd name of  | 
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));
 | 
|
| 6860 | 127  | 
|
128  | 
fun terminator false = Scan.succeed ()  | 
|
129  | 
| terminator true = P.group "terminator" (Scan.option P.sync -- P.$$$ ";" >> K ());  | 
|
130  | 
||
131  | 
in  | 
|
| 5829 | 132  | 
|
| 6860 | 133  | 
fun command term cmd =  | 
134  | 
P.$$$ ";" >> K None ||  | 
|
135  | 
P.sync >> K None ||  | 
|
| 7026 | 136  | 
(P.position P.command :-- command_body cmd) --| terminator term  | 
| 6860 | 137  | 
>> (fn ((name, pos), (int_only, f)) =>  | 
138  | 
Some (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>  | 
|
139  | 
Toplevel.interactive int_only |> f));  | 
|
| 5829 | 140  | 
|
| 6199 | 141  | 
end;  | 
142  | 
||
| 5829 | 143  | 
|
144  | 
||
145  | 
(** global syntax state **)  | 
|
146  | 
||
| 7026 | 147  | 
local  | 
148  | 
||
149  | 
val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon);  | 
|
| 7750 | 150  | 
val global_parsers =  | 
| 8660 | 151  | 
ref (Symtab.empty: (((string * string) * (bool * parser_fn)) * markup_kind option) Symtab.table);  | 
152  | 
val global_markups = ref ([]: (string * markup_kind) list);  | 
|
| 5952 | 153  | 
|
| 7026 | 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))
 | 
|
159  | 
end;  | 
|
| 5829 | 160  | 
|
| 7789 | 161  | 
fun get_markup (ms, (name, (_, Some m))) = (name, m) :: ms  | 
162  | 
| get_markup (ms, _) = ms;  | 
|
| 7750 | 163  | 
|
164  | 
fun make_markups () = global_markups := Symtab.foldl get_markup ([], ! global_parsers);  | 
|
165  | 
fun change_parsers f = (global_parsers := f (! global_parsers); make_markups ());  | 
|
| 6722 | 166  | 
|
| 7026 | 167  | 
in  | 
168  | 
||
| 7750 | 169  | 
|
170  | 
(* get current syntax *)  | 
|
| 7026 | 171  | 
|
172  | 
(*Note: the syntax for files is statically determined at the very  | 
|
173  | 
beginning; for interactive processing it may change dynamically.*)  | 
|
174  | 
||
175  | 
fun get_lexicons () = ! global_lexicons;  | 
|
176  | 
fun get_parsers () = ! global_parsers;  | 
|
| 7750 | 177  | 
fun get_parser () = apsome (#2 o #1) o curry Symtab.lookup (! global_parsers);  | 
| 7789 | 178  | 
|
179  | 
fun lookup_markup name = assoc (! global_markups, name);  | 
|
| 8660 | 180  | 
fun is_markup kind name = (case lookup_markup name of Some k => k = kind | None => false);  | 
| 5829 | 181  | 
|
182  | 
||
183  | 
(* augment syntax *)  | 
|
184  | 
||
| 7026 | 185  | 
fun add_keywords keywords = change_lexicons (apfst (fn lex =>  | 
186  | 
(Scan.extend_lexicon lex (map Symbol.explode keywords))));  | 
|
| 5829 | 187  | 
|
| 7750 | 188  | 
fun add_parser (tab, Parser (name, (comment, kind, markup), int_only, parse)) =  | 
| 5829 | 189  | 
(if is_none (Symtab.lookup (tab, name)) then ()  | 
190  | 
  else warning ("Redefined outer syntax command " ^ quote name);
 | 
|
| 7750 | 191  | 
Symtab.update ((name, (((comment, kind), (int_only, parse)), markup)), tab));  | 
| 5829 | 192  | 
|
193  | 
fun add_parsers parsers =  | 
|
| 7026 | 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))));  | 
|
197  | 
||
198  | 
end;  | 
|
| 5829 | 199  | 
|
200  | 
||
| 7026 | 201  | 
(* print syntax *)  | 
202  | 
||
203  | 
fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ()));  | 
|
204  | 
||
205  | 
fun dest_parsers () =  | 
|
| 7750 | 206  | 
map (fn (name, (((cmt, kind), (int_only, _)), _)) => (name, cmt, kind, int_only))  | 
| 7026 | 207  | 
(Symtab.dest (get_parsers ()));  | 
| 5829 | 208  | 
|
| 8720 | 209  | 
fun help_outer_syntax () =  | 
| 7026 | 210  | 
let  | 
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 ());  | 
|
214  | 
in  | 
|
| 8720 | 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)]  | 
|
| 7026 | 218  | 
end;  | 
| 5829 | 219  | 
|
| 8720 | 220  | 
val print_outer_syntax = Pretty.writeln o Pretty.chunks o help_outer_syntax;  | 
221  | 
||
| 7367 | 222  | 
val print_help =  | 
223  | 
Toplevel.keep (fn state =>  | 
|
| 7613 | 224  | 
let val opt_thy = try Toplevel.theory_of state in  | 
| 8720 | 225  | 
help_outer_syntax () @  | 
226  | 
Method.help_methods opt_thy @  | 
|
| 7613 | 227  | 
Attrib.help_attributes opt_thy  | 
| 8720 | 228  | 
|> Pretty.chunks |> Pretty.writeln  | 
| 7613 | 229  | 
end);  | 
| 7367 | 230  | 
|
| 5829 | 231  | 
|
232  | 
||
233  | 
(** read theory **)  | 
|
234  | 
||
| 7746 | 235  | 
(* special keywords *)  | 
| 6247 | 236  | 
|
| 7746 | 237  | 
val headerN = "header";  | 
| 6247 | 238  | 
val theoryN = "theory";  | 
| 7746 | 239  | 
|
| 7750 | 240  | 
val theory_keyword = P.$$$ theoryN;  | 
241  | 
val header_keyword = P.$$$ headerN;  | 
|
| 7746 | 242  | 
val semicolon = P.$$$ ";";  | 
| 6247 | 243  | 
|
244  | 
||
| 7683 | 245  | 
(* sources *)  | 
| 5829 | 246  | 
|
| 6860 | 247  | 
local  | 
| 5829 | 248  | 
|
| 6860 | 249  | 
val no_terminator =  | 
| 7750 | 250  | 
Scan.unless semicolon (Scan.one (T.not_sync andf T.not_eof));  | 
| 6860 | 251  | 
|
252  | 
val recover = Scan.prompt "recover# " (Scan.repeat no_terminator);  | 
|
253  | 
||
254  | 
in  | 
|
| 5829 | 255  | 
|
| 6860 | 256  | 
fun source term do_recover cmd src =  | 
| 5829 | 257  | 
src  | 
| 7750 | 258  | 
|> Source.source T.stopper  | 
259  | 
(Scan.bulk (fn xs => P.!!! (command term (cmd ())) xs))  | 
|
| 6860 | 260  | 
(if do_recover then Some recover else None)  | 
| 5829 | 261  | 
|> Source.mapfilter I;  | 
262  | 
||
| 6860 | 263  | 
end;  | 
264  | 
||
| 7683 | 265  | 
fun token_source (src, pos) =  | 
266  | 
src  | 
|
267  | 
|> Symbol.source false  | 
|
| 7750 | 268  | 
|> T.source false (K (get_lexicons ())) pos;  | 
| 7683 | 269  | 
|
270  | 
fun filter_proper src =  | 
|
271  | 
src  | 
|
| 7750 | 272  | 
|> Source.filter T.is_proper;  | 
| 7683 | 273  | 
|
| 5829 | 274  | 
|
| 7746 | 275  | 
(* scan header *)  | 
| 5829 | 276  | 
|
| 7026 | 277  | 
fun scan_header get_lex scan (src, pos) =  | 
| 5829 | 278  | 
src  | 
279  | 
|> Symbol.source false  | 
|
| 7750 | 280  | 
|> T.source false (fn () => (get_lex (), Scan.empty_lexicon)) pos  | 
| 7683 | 281  | 
|> filter_proper  | 
| 7750 | 282  | 
|> Source.source T.stopper (Scan.single scan) None  | 
| 5829 | 283  | 
|> (fst o the o Source.get_single);  | 
284  | 
||
| 7746 | 285  | 
|
286  | 
(* detect new/old header *)  | 
|
287  | 
||
288  | 
local  | 
|
| 5829 | 289  | 
|
| 7746 | 290  | 
val check_header_lexicon = Scan.make_lexicon [Symbol.explode headerN, Symbol.explode theoryN];  | 
291  | 
val check_header = Scan.option (header_keyword || theory_keyword);  | 
|
292  | 
||
293  | 
in  | 
|
294  | 
||
295  | 
fun is_old_theory src = is_none (scan_header (K check_header_lexicon) check_header src);  | 
|
296  | 
||
297  | 
end;  | 
|
| 6199 | 298  | 
|
299  | 
||
300  | 
(* deps_thy --- inspect theory header *)  | 
|
301  | 
||
| 7746 | 302  | 
local  | 
303  | 
||
| 6247 | 304  | 
val header_lexicon =  | 
| 7746 | 305  | 
  Scan.make_lexicon (map Symbol.explode ["(", ")", "+", ":", ";", "=", "files", headerN, theoryN]);
 | 
| 6199 | 306  | 
|
| 7026 | 307  | 
val file_name =  | 
308  | 
  (P.$$$ "(" |-- P.!!! (P.name --| P.$$$ ")")) >> rpair false || P.name >> rpair true;
 | 
|
| 6247 | 309  | 
|
| 7746 | 310  | 
in  | 
311  | 
||
312  | 
val theory_header =  | 
|
| 6860 | 313  | 
(P.name -- (P.$$$ "=" |-- P.enum1 "+" P.name) --  | 
| 7746 | 314  | 
Scan.optional (P.$$$ "files" |-- P.!!! (Scan.repeat1 file_name)) [] --| P.$$$ ":")  | 
| 6247 | 315  | 
>> (fn ((A, Bs), files) => (A, Bs, files));  | 
316  | 
||
| 7746 | 317  | 
val new_header =  | 
318  | 
header_keyword |-- (P.!!! (P.text -- Scan.option semicolon -- theory_keyword |-- theory_header))  | 
|
319  | 
|| theory_keyword |-- P.!!! theory_header;  | 
|
| 6199 | 320  | 
|
321  | 
val old_header =  | 
|
| 7810 | 322  | 
P.!!! (P.group "theory header"  | 
323  | 
(P.name -- (P.$$$ "=" |-- P.name -- Scan.repeat (P.$$$ "+" |-- P.name))))  | 
|
| 6247 | 324  | 
>> (fn (A, (B, Bs)) => (A, B :: Bs, []: (string * bool) list));  | 
| 6199 | 325  | 
|
| 7940 | 326  | 
fun deps_thy name ml path =  | 
| 6199 | 327  | 
let  | 
| 7735 | 328  | 
val src = Source.of_string (File.read path);  | 
329  | 
val pos = Path.position path;  | 
|
| 6247 | 330  | 
val (name', parents, files) =  | 
| 7940 | 331  | 
(*unfortunately, old-style headers dynamically depend on the current lexicon*)  | 
| 7746 | 332  | 
if is_old_theory (src, pos) then  | 
333  | 
scan_header ThySyn.get_lexicon (Scan.error old_header) (src, pos)  | 
|
| 7735 | 334  | 
else scan_header (K header_lexicon) (Scan.error new_header) (src, pos);  | 
| 6199 | 335  | 
|
336  | 
val ml_path = ThyLoad.ml_path name;  | 
|
| 7940 | 337  | 
val ml_file = if ml andalso is_some (ThyLoad.check_file ml_path) then [ml_path] else [];  | 
| 6199 | 338  | 
in  | 
339  | 
if name <> name' then  | 
|
| 7940 | 340  | 
      error ("Filename " ^ quote (Path.pack path) ^
 | 
| 8078 | 341  | 
" does not agree with theory name " ^ quote name')  | 
| 6247 | 342  | 
else (parents, map (Path.unpack o #1) files @ ml_file)  | 
| 6199 | 343  | 
end;  | 
344  | 
||
| 7746 | 345  | 
end;  | 
346  | 
||
| 6199 | 347  | 
|
| 7750 | 348  | 
(* present theory source *)  | 
349  | 
||
350  | 
local  | 
|
351  | 
||
| 8583 | 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;  | 
|
| 7903 | 354  | 
|
| 8583 | 355  | 
val improper_keep_indent = Scan.repeat  | 
356  | 
(Scan.unless (Scan.one T.is_indent -- Scan.one T.is_proper) (Scan.one is_improper));  | 
|
| 7903 | 357  | 
|
358  | 
val improper_end =  | 
|
359  | 
(improper -- semicolon) |-- improper_keep_indent ||  | 
|
360  | 
improper_keep_indent;  | 
|
361  | 
||
| 8583 | 362  | 
|
363  | 
val ignore =  | 
|
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)));  | 
|
367  | 
||
| 8652 | 368  | 
val opt_newline = Scan.option (Scan.one T.is_newline);  | 
369  | 
||
| 8583 | 370  | 
val ignore_stuff =  | 
| 8652 | 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);  | 
|
| 8583 | 373  | 
|
| 8660 | 374  | 
fun markup_kind k = Scan.one (T.is_kind T.Command andf is_markup k o T.val_of);  | 
375  | 
||
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;  | 
|
| 7750 | 379  | 
|
| 7755 | 380  | 
val present_token =  | 
| 8583 | 381  | 
ignore_stuff >> K None ||  | 
382  | 
(improper |-- markup -- P.!!!! (improper |-- P.text --| improper_end) >> Present.markup_token ||  | 
|
| 8660 | 383  | 
improper |-- markup_env -- P.!!!! (improper |-- P.text --| improper_end) >> Present.markup_env_token ||  | 
| 8583 | 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;  | 
|
| 7750 | 388  | 
|
389  | 
in  | 
|
390  | 
||
391  | 
(*note: lazy evaluation ahead*)  | 
|
392  | 
||
393  | 
fun present_toks text pos () =  | 
|
394  | 
token_source (Source.of_list (Library.untabify text), pos)  | 
|
| 8583 | 395  | 
|> Source.source T.stopper (Scan.bulk (Scan.error present_token)) None  | 
396  | 
|> Source.mapfilter I  | 
|
| 7755 | 397  | 
|> Source.exhaust;  | 
| 7750 | 398  | 
|
399  | 
fun present_text text () =  | 
|
400  | 
Source.exhaust (Symbol.source false (Source.of_list (Library.untabify text)));  | 
|
401  | 
||
402  | 
end;  | 
|
403  | 
||
404  | 
||
| 6199 | 405  | 
(* load_thy --- read text (including header) *)  | 
406  | 
||
| 7746 | 407  | 
local  | 
408  | 
||
| 7940 | 409  | 
fun try_ml_file name time =  | 
| 6199 | 410  | 
let  | 
411  | 
val path = ThyLoad.ml_path name;  | 
|
| 7940 | 412  | 
val tr = Toplevel.imperative (fn () => ThyInfo.load_file time path);  | 
| 6247 | 413  | 
val tr_name = if time then "time_use" else "use";  | 
| 6199 | 414  | 
in  | 
| 7243 | 415  | 
if is_none (ThyLoad.check_file path) then ()  | 
| 7062 | 416  | 
else Toplevel.excursion_error [Toplevel.empty |> Toplevel.name tr_name |> tr]  | 
| 6199 | 417  | 
end;  | 
418  | 
||
| 7683 | 419  | 
fun parse_thy src_pos =  | 
| 7746 | 420  | 
src_pos  | 
421  | 
|> token_source  | 
|
422  | 
|> filter_proper  | 
|
423  | 
|> source false false (K (get_parser ()))  | 
|
424  | 
|> Source.exhaust;  | 
|
| 5829 | 425  | 
|
| 6247 | 426  | 
fun run_thy name path =  | 
| 7683 | 427  | 
let  | 
| 7735 | 428  | 
val text = explode (File.read path);  | 
429  | 
val src = Source.of_list text;  | 
|
430  | 
val pos = Path.position path;  | 
|
| 7683 | 431  | 
in  | 
| 7735 | 432  | 
Present.init_theory name;  | 
| 7768 | 433  | 
Present.verbatim_source name (present_text text);  | 
| 8191 | 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!*)  | 
|
| 
7774
 
6da9b544a12d
Present.token_source after load (better errors!?);
 
wenzelm 
parents: 
7768 
diff
changeset
 | 
436  | 
else (Toplevel.excursion_error (parse_thy (src, pos));  | 
| 
 
6da9b544a12d
Present.token_source after load (better errors!?);
 
wenzelm 
parents: 
7768 
diff
changeset
 | 
437  | 
Present.token_source name (present_toks text pos))  | 
| 6247 | 438  | 
end;  | 
| 6199 | 439  | 
|
| 7746 | 440  | 
in  | 
441  | 
||
| 6199 | 442  | 
fun load_thy name ml time path =  | 
| 6247 | 443  | 
(if time then  | 
444  | 
timeit (fn () =>  | 
|
445  | 
     (writeln ("\n**** Starting theory " ^ quote name ^ " ****");
 | 
|
446  | 
setmp Goals.proof_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);  | 
|
| 7940 | 450  | 
if ml then try_ml_file name time else ());  | 
| 5829 | 451  | 
|
| 7746 | 452  | 
end;  | 
453  | 
||
| 5829 | 454  | 
|
455  | 
(* interactive source of state transformers *)  | 
|
456  | 
||
| 7333 | 457  | 
fun isar term no_pos =  | 
| 5829 | 458  | 
Source.tty  | 
459  | 
|> Symbol.source true  | 
|
| 7750 | 460  | 
|> T.source true get_lexicons  | 
| 7604 | 461  | 
(if no_pos then Position.none else Position.line_name 1 "stdin")  | 
| 7683 | 462  | 
|> filter_proper  | 
| 6860 | 463  | 
|> source term true get_parser;  | 
| 5829 | 464  | 
|
465  | 
||
466  | 
||
467  | 
(** the read-eval-print loop **)  | 
|
468  | 
||
| 5923 | 469  | 
(* main loop *)  | 
470  | 
||
| 7333 | 471  | 
fun gen_loop term no_pos =  | 
472  | 
(Context.reset_context ();  | 
|
473  | 
Toplevel.loop (isar term no_pos));  | 
|
| 5829 | 474  | 
|
| 7333 | 475  | 
fun gen_main term no_pos =  | 
| 5829 | 476  | 
(Toplevel.set_state Toplevel.toplevel;  | 
| 6199 | 477  | 
writeln (Session.welcome ());  | 
| 7333 | 478  | 
gen_loop term no_pos);  | 
| 6860 | 479  | 
|
| 7333 | 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;  | 
|
| 5829 | 484  | 
|
485  | 
||
486  | 
(* help *)  | 
|
487  | 
||
488  | 
fun help () =  | 
|
489  | 
  writeln ("This is Isabelle's underlying ML system (" ^ ml_system ^ ");\n\
 | 
|
| 5883 | 490  | 
\invoke 'loop();' to enter the Isar loop.");  | 
| 5829 | 491  | 
|
492  | 
||
| 6373 | 493  | 
(*final declarations of this structure!*)  | 
| 7789 | 494  | 
val command = parser false None;  | 
| 8660 | 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);  | 
|
| 7789 | 498  | 
val improper_command = parser true None;  | 
| 6685 | 499  | 
|
| 6373 | 500  | 
|
| 5829 | 501  | 
end;  | 
502  | 
||
| 6199 | 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;  | 
|
507  | 
||
| 5829 | 508  | 
structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax;  | 
509  | 
open BasicOuterSyntax;  |