| author | wenzelm | 
| Mon, 17 May 2010 15:11:25 +0200 | |
| changeset 36959 | f5417836dbea | 
| parent 36955 | 226fb165833e | 
| child 37216 | 3165bc303f66 | 
| permissions | -rw-r--r-- | 
| 5829 | 1  | 
(* Title: Pure/Isar/outer_syntax.ML  | 
2  | 
Author: Markus Wenzel, TU Muenchen  | 
|
3  | 
||
| 
27353
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
26990 
diff
changeset
 | 
4  | 
The global Isabelle/Isar outer syntax.  | 
| 
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
26990 
diff
changeset
 | 
5  | 
|
| 
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
26990 
diff
changeset
 | 
6  | 
Note: the syntax for files is statically determined at the very  | 
| 
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
26990 
diff
changeset
 | 
7  | 
beginning; for interactive processing it may change dynamically.  | 
| 5829 | 8  | 
*)  | 
9  | 
||
10  | 
signature OUTER_SYNTAX =  | 
|
11  | 
sig  | 
|
| 36950 | 12  | 
val command: string -> string -> Keyword.T ->  | 
| 29311 | 13  | 
(Toplevel.transition -> Toplevel.transition) parser -> unit  | 
| 36950 | 14  | 
val markup_command: ThyOutput.markup -> string -> string -> Keyword.T ->  | 
| 29311 | 15  | 
(Toplevel.transition -> Toplevel.transition) parser -> unit  | 
| 36950 | 16  | 
val improper_command: string -> string -> Keyword.T ->  | 
| 29311 | 17  | 
(Toplevel.transition -> Toplevel.transition) parser -> unit  | 
18  | 
val internal_command: string -> (Toplevel.transition -> Toplevel.transition) parser -> unit  | 
|
| 36950 | 19  | 
val local_theory': string -> string -> Keyword.T ->  | 
| 29380 | 20  | 
(bool -> local_theory -> local_theory) parser -> unit  | 
| 36950 | 21  | 
val local_theory: string -> string -> Keyword.T ->  | 
| 29311 | 22  | 
(local_theory -> local_theory) parser -> unit  | 
| 36950 | 23  | 
val local_theory_to_proof': string -> string -> Keyword.T ->  | 
| 29311 | 24  | 
(bool -> local_theory -> Proof.state) parser -> unit  | 
| 36950 | 25  | 
val local_theory_to_proof: string -> string -> Keyword.T ->  | 
| 29311 | 26  | 
(local_theory -> Proof.state) parser -> unit  | 
| 5883 | 27  | 
val print_outer_syntax: unit -> unit  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
28  | 
val scan: Position.T -> string -> Token.T list  | 
| 25580 | 29  | 
val parse: Position.T -> string -> Toplevel.transition list  | 
| 26431 | 30  | 
val process_file: Path.T -> theory -> theory  | 
| 26600 | 31  | 
type isar  | 
32  | 
val isar: bool -> isar  | 
|
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
33  | 
val prepare_command: Position.T -> string -> Toplevel.transition  | 
| 
29428
 
3ab54b42ded8
load_thy: explicit after_load phase for presentation;
 
wenzelm 
parents: 
29380 
diff
changeset
 | 
34  | 
val load_thy: string -> Position.T -> string list -> bool -> unit -> unit  | 
| 5829 | 35  | 
end;  | 
36  | 
||
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
37  | 
structure Outer_Syntax: OUTER_SYNTAX =  | 
| 5829 | 38  | 
struct  | 
39  | 
||
40  | 
(** outer syntax **)  | 
|
41  | 
||
| 29311 | 42  | 
(* command parsers *)  | 
| 5829 | 43  | 
|
| 29311 | 44  | 
datatype command = Command of  | 
| 24868 | 45  | 
 {comment: string,
 | 
46  | 
markup: ThyOutput.markup option,  | 
|
47  | 
int_only: bool,  | 
|
| 29311 | 48  | 
parse: (Toplevel.transition -> Toplevel.transition) parser};  | 
| 5829 | 49  | 
|
| 29311 | 50  | 
fun make_command comment markup int_only parse =  | 
51  | 
  Command {comment = comment, markup = markup, int_only = int_only, parse = parse};
 | 
|
| 5829 | 52  | 
|
53  | 
||
54  | 
(* parse command *)  | 
|
55  | 
||
| 6860 | 56  | 
local  | 
| 6199 | 57  | 
|
| 
14925
 
0f86a8a694f8
added read (provides transition names and sources);
 
wenzelm 
parents: 
14687 
diff
changeset
 | 
58  | 
fun terminate false = Scan.succeed ()  | 
| 36950 | 59  | 
| terminate true = Parse.group "end of input" (Scan.option Parse.sync -- Parse.semicolon >> K ());  | 
| 
14925
 
0f86a8a694f8
added read (provides transition names and sources);
 
wenzelm 
parents: 
14687 
diff
changeset
 | 
60  | 
|
| 26620 | 61  | 
fun body cmd (name, _) =  | 
| 7026 | 62  | 
(case cmd name of  | 
| 29311 | 63  | 
    SOME (Command {int_only, parse, ...}) =>
 | 
| 36950 | 64  | 
Parse.!!! (Scan.prompt (name ^ "# ") (Parse.tags |-- parse >> pair int_only))  | 
| 15531 | 65  | 
  | NONE => sys_error ("no parser for outer syntax command " ^ quote name));
 | 
| 6860 | 66  | 
|
67  | 
in  | 
|
| 5829 | 68  | 
|
| 26620 | 69  | 
fun parse_command do_terminate cmd =  | 
| 36950 | 70  | 
Parse.semicolon >> K NONE ||  | 
71  | 
Parse.sync >> K NONE ||  | 
|
72  | 
(Parse.position Parse.command :-- body cmd) --| terminate do_terminate  | 
|
| 6860 | 73  | 
>> (fn ((name, pos), (int_only, f)) =>  | 
| 15531 | 74  | 
SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>  | 
| 6860 | 75  | 
Toplevel.interactive int_only |> f));  | 
| 5829 | 76  | 
|
| 6199 | 77  | 
end;  | 
78  | 
||
| 5829 | 79  | 
|
80  | 
||
| 9132 | 81  | 
(** global outer syntax **)  | 
| 5829 | 82  | 
|
| 7026 | 83  | 
local  | 
84  | 
||
| 32738 | 85  | 
val global_commands = Unsynchronized.ref (Symtab.empty: command Symtab.table);  | 
86  | 
val global_markups = Unsynchronized.ref ([]: (string * ThyOutput.markup) list);  | 
|
| 5952 | 87  | 
|
| 29311 | 88  | 
fun change_commands f = CRITICAL (fn () =>  | 
| 32738 | 89  | 
(Unsynchronized.change global_commands f;  | 
| 23939 | 90  | 
global_markups :=  | 
| 29311 | 91  | 
    Symtab.fold (fn (name, Command {markup = SOME m, ...}) => cons (name, m) | _ => I)
 | 
92  | 
(! global_commands) []));  | 
|
| 6722 | 93  | 
|
| 7026 | 94  | 
in  | 
95  | 
||
| 9132 | 96  | 
(* access current syntax *)  | 
| 7026 | 97  | 
|
| 33223 | 98  | 
fun get_commands () = ! global_commands;  | 
99  | 
fun get_markups () = ! global_markups;  | 
|
| 7026 | 100  | 
|
| 29311 | 101  | 
fun get_command () = Symtab.lookup (get_commands ());  | 
| 36950 | 102  | 
fun get_syntax () = CRITICAL (fn () => (Keyword.get_lexicons (), get_command ()));  | 
| 7789 | 103  | 
|
| 24868 | 104  | 
fun is_markup kind name = AList.lookup (op =) (get_markups ()) name = SOME kind;  | 
| 5829 | 105  | 
|
106  | 
||
107  | 
(* augment syntax *)  | 
|
108  | 
||
| 29311 | 109  | 
fun add_command name kind cmd = CRITICAL (fn () =>  | 
| 36950 | 110  | 
(Keyword.command name kind;  | 
| 29311 | 111  | 
if not (Symtab.defined (get_commands ()) name) then ()  | 
| 24868 | 112  | 
  else warning ("Redefining outer syntax command " ^ quote name);
 | 
| 29311 | 113  | 
change_commands (Symtab.update (name, cmd))));  | 
| 5829 | 114  | 
|
| 24868 | 115  | 
fun command name comment kind parse =  | 
| 29311 | 116  | 
add_command name kind (make_command comment NONE false parse);  | 
| 5829 | 117  | 
|
| 24868 | 118  | 
fun markup_command markup name comment kind parse =  | 
| 29311 | 119  | 
add_command name kind (make_command comment (SOME markup) false parse);  | 
| 24868 | 120  | 
|
121  | 
fun improper_command name comment kind parse =  | 
|
| 29311 | 122  | 
add_command name kind (make_command comment NONE true parse);  | 
| 7026 | 123  | 
|
124  | 
end;  | 
|
| 5829 | 125  | 
|
| 29311 | 126  | 
fun internal_command name parse =  | 
| 36950 | 127  | 
command name "(internal)" Keyword.control (parse >> (fn tr => Toplevel.no_timing o tr));  | 
| 29311 | 128  | 
|
| 5829 | 129  | 
|
| 26990 | 130  | 
(* local_theory commands *)  | 
131  | 
||
132  | 
fun local_theory_command do_print trans name comment kind parse =  | 
|
| 36950 | 133  | 
command name comment kind (Parse.opt_target -- parse  | 
| 26990 | 134  | 
>> (fn (loc, f) => (if do_print then Toplevel.print else I) o trans loc f));  | 
135  | 
||
| 29380 | 136  | 
val local_theory' = local_theory_command false Toplevel.local_theory';  | 
| 29311 | 137  | 
val local_theory = local_theory_command false Toplevel.local_theory;  | 
| 26990 | 138  | 
val local_theory_to_proof' = local_theory_command true Toplevel.local_theory_to_proof';  | 
| 29311 | 139  | 
val local_theory_to_proof = local_theory_command true Toplevel.local_theory_to_proof;  | 
| 26990 | 140  | 
|
141  | 
||
| 24872 | 142  | 
(* inspect syntax *)  | 
| 7026 | 143  | 
|
| 29311 | 144  | 
fun dest_commands () =  | 
145  | 
get_commands () |> Symtab.dest |> sort_wrt #1  | 
|
146  | 
  |> map (fn (name, Command {comment, int_only, ...}) => (name, comment, int_only));
 | 
|
| 5829 | 147  | 
|
| 9223 | 148  | 
fun print_outer_syntax () =  | 
| 7026 | 149  | 
let  | 
| 
27353
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
26990 
diff
changeset
 | 
150  | 
fun pretty_cmd (name, comment, _) =  | 
| 7026 | 151  | 
Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];  | 
| 29311 | 152  | 
val (int_cmds, cmds) = List.partition #3 (dest_commands ());  | 
| 7026 | 153  | 
in  | 
| 36950 | 154  | 
    [Pretty.strs ("syntax keywords:" :: map quote (Keyword.dest_keywords ())),
 | 
| 18326 | 155  | 
Pretty.big_list "commands:" (map pretty_cmd cmds),  | 
156  | 
Pretty.big_list "interactive-only commands:" (map pretty_cmd int_cmds)]  | 
|
| 9223 | 157  | 
|> Pretty.chunks |> Pretty.writeln  | 
| 7026 | 158  | 
end;  | 
| 5829 | 159  | 
|
160  | 
||
161  | 
||
| 9132 | 162  | 
(** toplevel parsing **)  | 
| 5829 | 163  | 
|
| 9132 | 164  | 
(* basic sources *)  | 
| 6860 | 165  | 
|
| 26620 | 166  | 
fun toplevel_source term do_recover cmd src =  | 
| 9132 | 167  | 
let  | 
168  | 
val no_terminator =  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
169  | 
Scan.unless Parse.semicolon (Scan.one (Token.not_sync andf Token.not_eof));  | 
| 
23682
 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 
wenzelm 
parents: 
23679 
diff
changeset
 | 
170  | 
fun recover int =  | 
| 
 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 
wenzelm 
parents: 
23679 
diff
changeset
 | 
171  | 
(int, fn _ => Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]);  | 
| 9132 | 172  | 
in  | 
173  | 
src  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
174  | 
|> Token.source_proper  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
175  | 
|> Source.source Token.stopper  | 
| 36950 | 176  | 
(Scan.bulk (Parse.$$$ "--" -- Parse.!!! Parse.doc_source >> K NONE || Parse.not_eof >> SOME))  | 
| 
23682
 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 
wenzelm 
parents: 
23679 
diff
changeset
 | 
177  | 
(Option.map recover do_recover)  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19060 
diff
changeset
 | 
178  | 
|> Source.map_filter I  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
179  | 
|> Source.source Token.stopper  | 
| 36950 | 180  | 
(Scan.bulk (fn xs => Parse.!!! (parse_command term (cmd ())) xs))  | 
| 
23682
 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 
wenzelm 
parents: 
23679 
diff
changeset
 | 
181  | 
(Option.map recover do_recover)  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19060 
diff
changeset
 | 
182  | 
|> Source.map_filter I  | 
| 9132 | 183  | 
end;  | 
| 5829 | 184  | 
|
| 7746 | 185  | 
|
| 25580 | 186  | 
(* off-line scanning/parsing *)  | 
| 
14925
 
0f86a8a694f8
added read (provides transition names and sources);
 
wenzelm 
parents: 
14687 
diff
changeset
 | 
187  | 
|
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
188  | 
fun scan pos str =  | 
| 16195 | 189  | 
Source.of_string str  | 
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
190  | 
  |> Symbol.source {do_recover = false}
 | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
191  | 
  |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
 | 
| 16195 | 192  | 
|> Source.exhaust;  | 
193  | 
||
| 25580 | 194  | 
fun parse pos str =  | 
195  | 
Source.of_string str  | 
|
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
196  | 
  |> Symbol.source {do_recover = false}
 | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
197  | 
  |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
 | 
| 29311 | 198  | 
|> toplevel_source false NONE get_command  | 
| 25580 | 199  | 
|> Source.exhaust;  | 
200  | 
||
| 14091 | 201  | 
|
| 26431 | 202  | 
(* process file *)  | 
203  | 
||
204  | 
fun process_file path thy =  | 
|
205  | 
let  | 
|
| 26881 | 206  | 
val trs = parse (Path.position path) (File.read path);  | 
| 
28424
 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 
wenzelm 
parents: 
27872 
diff
changeset
 | 
207  | 
val init = Toplevel.init_theory "" (K thy) (K ()) Toplevel.empty;  | 
| 
 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 
wenzelm 
parents: 
27872 
diff
changeset
 | 
208  | 
val result = fold Toplevel.command (init :: trs) Toplevel.toplevel;  | 
| 
 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 
wenzelm 
parents: 
27872 
diff
changeset
 | 
209  | 
in  | 
| 
 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 
wenzelm 
parents: 
27872 
diff
changeset
 | 
210  | 
(case (Toplevel.is_theory result, Toplevel.generic_theory_of result) of  | 
| 
 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 
wenzelm 
parents: 
27872 
diff
changeset
 | 
211  | 
(true, Context.Theory thy') => thy'  | 
| 
 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 
wenzelm 
parents: 
27872 
diff
changeset
 | 
212  | 
| _ => error "Bad result state: global theory expected")  | 
| 
 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 
wenzelm 
parents: 
27872 
diff
changeset
 | 
213  | 
end;  | 
| 26431 | 214  | 
|
215  | 
||
| 24868 | 216  | 
(* interactive source of toplevel transformers *)  | 
217  | 
||
| 26600 | 218  | 
type isar =  | 
219  | 
(Toplevel.transition, (Toplevel.transition option,  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
220  | 
(Token.T, (Token.T option, (Token.T, (Token.T,  | 
| 30573 | 221  | 
(Symbol_Pos.T, Position.T * (Symbol.symbol, (string, unit) Source.source)  | 
| 27770 | 222  | 
Source.source) Source.source) Source.source) Source.source)  | 
223  | 
Source.source) Source.source) Source.source) Source.source;  | 
|
| 26600 | 224  | 
|
225  | 
fun isar term : isar =  | 
|
| 24868 | 226  | 
Source.tty  | 
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
227  | 
  |> Symbol.source {do_recover = true}
 | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
228  | 
  |> Token.source {do_recover = SOME true} Keyword.get_lexicons Position.none
 | 
| 29311 | 229  | 
|> toplevel_source term (SOME true) get_command;  | 
| 24868 | 230  | 
|
231  | 
||
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
232  | 
(* prepare toplevel commands -- fail-safe *)  | 
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
233  | 
|
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
234  | 
val not_singleton = "Exactly one command expected";  | 
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
235  | 
|
| 29311 | 236  | 
fun prepare_span commands span =  | 
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
237  | 
let  | 
| 
29315
 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 
wenzelm 
parents: 
29311 
diff
changeset
 | 
238  | 
val range_pos = Position.encode_range (ThySyntax.span_range span);  | 
| 
 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 
wenzelm 
parents: 
29311 
diff
changeset
 | 
239  | 
val toks = ThySyntax.span_content span;  | 
| 
 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 
wenzelm 
parents: 
29311 
diff
changeset
 | 
240  | 
val _ = List.app ThySyntax.report_token toks;  | 
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
241  | 
in  | 
| 29311 | 242  | 
(case Source.exhaust (toplevel_source false NONE (K commands) (Source.of_list toks)) of  | 
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
243  | 
[tr] => (tr, true)  | 
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
244  | 
| [] => (Toplevel.ignored range_pos, false)  | 
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
245  | 
| _ => (Toplevel.malformed range_pos not_singleton, true))  | 
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
246  | 
handle ERROR msg => (Toplevel.malformed range_pos msg, true)  | 
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
247  | 
end;  | 
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
248  | 
|
| 29311 | 249  | 
fun prepare_unit commands (cmd, proof, proper_proof) =  | 
| 
28436
 
4faf705a177d
load_thy: more precise treatment of improper cmd or proof (notably 'oops');
 
wenzelm 
parents: 
28432 
diff
changeset
 | 
250  | 
let  | 
| 29311 | 251  | 
val (tr, proper_cmd) = prepare_span commands cmd;  | 
252  | 
val proof_trs = map (prepare_span commands) proof |> filter #2 |> map #1;  | 
|
| 
28436
 
4faf705a177d
load_thy: more precise treatment of improper cmd or proof (notably 'oops');
 
wenzelm 
parents: 
28432 
diff
changeset
 | 
253  | 
in  | 
| 
 
4faf705a177d
load_thy: more precise treatment of improper cmd or proof (notably 'oops');
 
wenzelm 
parents: 
28432 
diff
changeset
 | 
254  | 
if proper_cmd andalso proper_proof then [(tr, proof_trs)]  | 
| 
 
4faf705a177d
load_thy: more precise treatment of improper cmd or proof (notably 'oops');
 
wenzelm 
parents: 
28432 
diff
changeset
 | 
255  | 
else map (rpair []) (if proper_cmd then tr :: proof_trs else proof_trs)  | 
| 
 
4faf705a177d
load_thy: more precise treatment of improper cmd or proof (notably 'oops');
 
wenzelm 
parents: 
28432 
diff
changeset
 | 
256  | 
end;  | 
| 
28432
 
944cb67f809a
load_thy: Toplevel.excursion based on units of command/proof pairs;
 
wenzelm 
parents: 
28424 
diff
changeset
 | 
257  | 
|
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
258  | 
fun prepare_command pos str =  | 
| 29311 | 259  | 
let val (lexs, commands) = get_syntax () in  | 
| 
29315
 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 
wenzelm 
parents: 
29311 
diff
changeset
 | 
260  | 
(case ThySyntax.parse_spans lexs pos str of  | 
| 29311 | 261  | 
[span] => #1 (prepare_span commands span)  | 
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
262  | 
| _ => Toplevel.malformed pos not_singleton)  | 
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
263  | 
end;  | 
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
264  | 
|
| 
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
265  | 
|
| 26611 | 266  | 
(* load_thy *)  | 
| 7746 | 267  | 
|
| 
27839
 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
268  | 
fun load_thy name pos text time =  | 
| 7683 | 269  | 
let  | 
| 29311 | 270  | 
val (lexs, commands) = get_syntax ();  | 
| 
23866
 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 
wenzelm 
parents: 
23796 
diff
changeset
 | 
271  | 
|
| 17932 | 272  | 
val _ = Present.init_theory name;  | 
| 
27855
 
b1bf607e06c2
load_thy: no untabify (preserve position information!), present spans instead of verbatim source;
 
wenzelm 
parents: 
27839 
diff
changeset
 | 
273  | 
|
| 
29315
 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 
wenzelm 
parents: 
29311 
diff
changeset
 | 
274  | 
val toks = Source.exhausted (ThySyntax.token_source lexs pos (Source.of_list text));  | 
| 
 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 
wenzelm 
parents: 
29311 
diff
changeset
 | 
275  | 
val spans = Source.exhaust (ThySyntax.span_source toks);  | 
| 
 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 
wenzelm 
parents: 
29311 
diff
changeset
 | 
276  | 
val _ = List.app ThySyntax.report_span spans;  | 
| 
 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 
wenzelm 
parents: 
29311 
diff
changeset
 | 
277  | 
val units = Source.exhaust (ThySyntax.unit_source (Source.of_list spans))  | 
| 29311 | 278  | 
|> maps (prepare_unit commands);  | 
| 
23866
 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 
wenzelm 
parents: 
23796 
diff
changeset
 | 
279  | 
|
| 
27855
 
b1bf607e06c2
load_thy: no untabify (preserve position information!), present spans instead of verbatim source;
 
wenzelm 
parents: 
27839 
diff
changeset
 | 
280  | 
val _ = Present.theory_source name  | 
| 
29315
 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 
wenzelm 
parents: 
29311 
diff
changeset
 | 
281  | 
(fn () => HTML.html_mode (implode o map ThySyntax.present_span) spans);  | 
| 
27855
 
b1bf607e06c2
load_thy: no untabify (preserve position information!), present spans instead of verbatim source;
 
wenzelm 
parents: 
27839 
diff
changeset
 | 
282  | 
|
| 
23866
 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 
wenzelm 
parents: 
23796 
diff
changeset
 | 
283  | 
    val _ = if time then writeln ("\n**** Starting theory " ^ quote name ^ " ****") else ();
 | 
| 
29428
 
3ab54b42ded8
load_thy: explicit after_load phase for presentation;
 
wenzelm 
parents: 
29380 
diff
changeset
 | 
284  | 
val results = cond_timeit time "" (fn () => Toplevel.excursion units);  | 
| 
23866
 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 
wenzelm 
parents: 
23796 
diff
changeset
 | 
285  | 
    val _ = if time then writeln ("**** Finished theory " ^ quote name ^ " ****\n") else ();
 | 
| 
29428
 
3ab54b42ded8
load_thy: explicit after_load phase for presentation;
 
wenzelm 
parents: 
29380 
diff
changeset
 | 
286  | 
|
| 
 
3ab54b42ded8
load_thy: explicit after_load phase for presentation;
 
wenzelm 
parents: 
29380 
diff
changeset
 | 
287  | 
fun after_load () =  | 
| 36950 | 288  | 
ThyOutput.present_thy (#1 lexs) Keyword.command_tags is_markup (Lazy.force results) toks  | 
| 
29428
 
3ab54b42ded8
load_thy: explicit after_load phase for presentation;
 
wenzelm 
parents: 
29380 
diff
changeset
 | 
289  | 
|> Buffer.content  | 
| 
 
3ab54b42ded8
load_thy: explicit after_load phase for presentation;
 
wenzelm 
parents: 
29380 
diff
changeset
 | 
290  | 
|> Present.theory_output name;  | 
| 
 
3ab54b42ded8
load_thy: explicit after_load phase for presentation;
 
wenzelm 
parents: 
29380 
diff
changeset
 | 
291  | 
in after_load end;  | 
| 
23866
 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 
wenzelm 
parents: 
23796 
diff
changeset
 | 
292  | 
|
| 5829 | 293  | 
end;  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
294  |