| author | wenzelm | 
| Thu, 01 Dec 2011 20:54:48 +0100 | |
| changeset 45717 | b4e7b9968e60 | 
| parent 44736 | c2a3f1c84179 | 
| child 46876 | 8f3bb485f628 | 
| 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: 
26990diff
changeset | 4 | The global Isabelle/Isar outer syntax. | 
| 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 5 | |
| 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
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: 
26990diff
changeset | 7 | beginning; for interactive processing it may change dynamically. | 
| 5829 | 8 | *) | 
| 9 | ||
| 10 | signature OUTER_SYNTAX = | |
| 11 | sig | |
| 43711 | 12 | type outer_syntax | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43711diff
changeset | 13 | val is_markup: outer_syntax -> Thy_Output.markup -> string -> bool | 
| 43711 | 14 | val get_syntax: unit -> (Scan.lexicon * Scan.lexicon) * outer_syntax | 
| 36950 | 15 | val command: string -> string -> Keyword.T -> | 
| 29311 | 16 | (Toplevel.transition -> Toplevel.transition) parser -> unit | 
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36959diff
changeset | 17 | val markup_command: Thy_Output.markup -> string -> string -> Keyword.T -> | 
| 29311 | 18 | (Toplevel.transition -> Toplevel.transition) parser -> unit | 
| 36950 | 19 | val improper_command: string -> string -> Keyword.T -> | 
| 29311 | 20 | (Toplevel.transition -> Toplevel.transition) parser -> unit | 
| 43711 | 21 | val internal_command: string -> | 
| 22 | (Toplevel.transition -> Toplevel.transition) parser -> unit | |
| 36950 | 23 | val local_theory': string -> string -> Keyword.T -> | 
| 29380 | 24 | (bool -> local_theory -> local_theory) parser -> unit | 
| 36950 | 25 | val local_theory: string -> string -> Keyword.T -> | 
| 29311 | 26 | (local_theory -> local_theory) parser -> unit | 
| 36950 | 27 | val local_theory_to_proof': string -> string -> Keyword.T -> | 
| 29311 | 28 | (bool -> local_theory -> Proof.state) parser -> unit | 
| 36950 | 29 | val local_theory_to_proof: string -> string -> Keyword.T -> | 
| 29311 | 30 | (local_theory -> Proof.state) parser -> unit | 
| 5883 | 31 | 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: 
36955diff
changeset | 32 | val scan: Position.T -> string -> Token.T list | 
| 25580 | 33 | val parse: Position.T -> string -> Toplevel.transition list | 
| 26431 | 34 | val process_file: Path.T -> theory -> theory | 
| 26600 | 35 | type isar | 
| 38253 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
38236diff
changeset | 36 | val isar: TextIO.instream -> bool -> isar | 
| 44659 | 37 | val read_command: Position.T -> string -> Toplevel.transition | 
| 44478 
4fdb1009a370
tuned signature -- emphasize traditional read/eval/print terminology, which is still relevant here;
 wenzelm parents: 
44357diff
changeset | 38 | val read_element: outer_syntax -> (unit -> theory) -> Thy_Syntax.element -> | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43711diff
changeset | 39 | (Toplevel.transition * Toplevel.transition list) list | 
| 5829 | 40 | end; | 
| 41 | ||
| 36953 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 wenzelm parents: 
36950diff
changeset | 42 | structure Outer_Syntax: OUTER_SYNTAX = | 
| 5829 | 43 | struct | 
| 44 | ||
| 45 | (** outer syntax **) | |
| 46 | ||
| 29311 | 47 | (* command parsers *) | 
| 5829 | 48 | |
| 29311 | 49 | datatype command = Command of | 
| 24868 | 50 |  {comment: string,
 | 
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36959diff
changeset | 51 | markup: Thy_Output.markup option, | 
| 24868 | 52 | int_only: bool, | 
| 29311 | 53 | parse: (Toplevel.transition -> Toplevel.transition) parser}; | 
| 5829 | 54 | |
| 29311 | 55 | fun make_command comment markup int_only parse = | 
| 56 |   Command {comment = comment, markup = markup, int_only = int_only, parse = parse};
 | |
| 5829 | 57 | |
| 58 | ||
| 59 | (* parse command *) | |
| 60 | ||
| 6860 | 61 | local | 
| 6199 | 62 | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 63 | fun terminate false = Scan.succeed () | 
| 44357 | 64 | | terminate true = | 
| 65 | Parse.group (fn () => "end of input") | |
| 66 | (Scan.option Parse.sync -- Parse.semicolon >> K ()); | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 67 | |
| 26620 | 68 | fun body cmd (name, _) = | 
| 7026 | 69 | (case cmd name of | 
| 29311 | 70 |     SOME (Command {int_only, parse, ...}) =>
 | 
| 36950 | 71 | Parse.!!! (Scan.prompt (name ^ "# ") (Parse.tags |-- parse >> pair int_only)) | 
| 37852 
a902f158b4fc
eliminated old-style sys_error/SYS_ERROR in favour of exception Fail -- after careful checking that there is no overlap with existing handling of that;
 wenzelm parents: 
37713diff
changeset | 72 |   | NONE => raise Fail ("No parser for outer syntax command " ^ quote name));
 | 
| 6860 | 73 | |
| 74 | in | |
| 5829 | 75 | |
| 26620 | 76 | fun parse_command do_terminate cmd = | 
| 36950 | 77 | Parse.semicolon >> K NONE || | 
| 78 | Parse.sync >> K NONE || | |
| 79 | (Parse.position Parse.command :-- body cmd) --| terminate do_terminate | |
| 6860 | 80 | >> (fn ((name, pos), (int_only, f)) => | 
| 15531 | 81 | SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |> | 
| 6860 | 82 | Toplevel.interactive int_only |> f)); | 
| 5829 | 83 | |
| 6199 | 84 | end; | 
| 85 | ||
| 5829 | 86 | |
| 43711 | 87 | (* type outer_syntax *) | 
| 88 | ||
| 89 | datatype outer_syntax = Outer_Syntax of | |
| 90 |  {commands: command Symtab.table,
 | |
| 91 | markups: (string * Thy_Output.markup) list}; | |
| 92 | ||
| 93 | fun make_outer_syntax commands markups = | |
| 94 |   Outer_Syntax {commands = commands, markups = markups};
 | |
| 95 | ||
| 96 | val empty_outer_syntax = make_outer_syntax Symtab.empty []; | |
| 97 | ||
| 98 | ||
| 99 | fun map_commands f (Outer_Syntax {commands, ...}) =
 | |
| 100 | let | |
| 101 | val commands' = f commands; | |
| 102 | val markups' = | |
| 103 |       Symtab.fold (fn (name, Command {markup = SOME m, ...}) => cons (name, m) | _ => I)
 | |
| 104 | commands' []; | |
| 105 | in make_outer_syntax commands' markups' end; | |
| 106 | ||
| 107 | fun dest_commands (Outer_Syntax {commands, ...}) =
 | |
| 108 | commands |> Symtab.dest |> sort_wrt #1 | |
| 109 |   |> map (fn (name, Command {comment, int_only, ...}) => (name, comment, int_only));
 | |
| 110 | ||
| 111 | fun lookup_commands (Outer_Syntax {commands, ...}) = Symtab.lookup commands;
 | |
| 112 | ||
| 113 | fun is_markup (Outer_Syntax {markups, ...}) kind name =
 | |
| 114 | AList.lookup (op =) markups name = SOME kind; | |
| 115 | ||
| 116 | ||
| 5829 | 117 | |
| 9132 | 118 | (** global outer syntax **) | 
| 5829 | 119 | |
| 7026 | 120 | local | 
| 121 | ||
| 43711 | 122 | (*synchronized wrt. Keywords*) | 
| 123 | val global_outer_syntax = Unsynchronized.ref empty_outer_syntax; | |
| 5952 | 124 | |
| 43711 | 125 | fun add_command name kind cmd = CRITICAL (fn () => | 
| 126 | (Keyword.command name kind; | |
| 127 | Unsynchronized.change global_outer_syntax (map_commands (fn commands => | |
| 128 | (if not (Symtab.defined commands name) then () | |
| 129 |     else warning ("Redefining outer syntax command " ^ quote name);
 | |
| 130 | Symtab.update (name, cmd) commands))))); | |
| 6722 | 131 | |
| 7026 | 132 | in | 
| 133 | ||
| 43711 | 134 | fun get_syntax () = CRITICAL (fn () => (Keyword.get_lexicons (), ! global_outer_syntax)); | 
| 7789 | 135 | |
| 43711 | 136 | fun lookup_commands_dynamic () = lookup_commands (! global_outer_syntax); | 
| 5829 | 137 | |
| 24868 | 138 | fun command name comment kind parse = | 
| 29311 | 139 | add_command name kind (make_command comment NONE false parse); | 
| 5829 | 140 | |
| 24868 | 141 | fun markup_command markup name comment kind parse = | 
| 29311 | 142 | add_command name kind (make_command comment (SOME markup) false parse); | 
| 24868 | 143 | |
| 144 | fun improper_command name comment kind parse = | |
| 29311 | 145 | add_command name kind (make_command comment NONE true parse); | 
| 7026 | 146 | |
| 29311 | 147 | fun internal_command name parse = | 
| 36950 | 148 | command name "(internal)" Keyword.control (parse >> (fn tr => Toplevel.no_timing o tr)); | 
| 29311 | 149 | |
| 43711 | 150 | end; | 
| 151 | ||
| 5829 | 152 | |
| 26990 | 153 | (* local_theory commands *) | 
| 154 | ||
| 155 | fun local_theory_command do_print trans name comment kind parse = | |
| 36950 | 156 | command name comment kind (Parse.opt_target -- parse | 
| 26990 | 157 | >> (fn (loc, f) => (if do_print then Toplevel.print else I) o trans loc f)); | 
| 158 | ||
| 29380 | 159 | val local_theory' = local_theory_command false Toplevel.local_theory'; | 
| 29311 | 160 | val local_theory = local_theory_command false Toplevel.local_theory; | 
| 26990 | 161 | val local_theory_to_proof' = local_theory_command true Toplevel.local_theory_to_proof'; | 
| 29311 | 162 | val local_theory_to_proof = local_theory_command true Toplevel.local_theory_to_proof; | 
| 26990 | 163 | |
| 164 | ||
| 24872 | 165 | (* inspect syntax *) | 
| 7026 | 166 | |
| 9223 | 167 | fun print_outer_syntax () = | 
| 7026 | 168 | let | 
| 43711 | 169 | val (keywords, outer_syntax) = | 
| 170 | CRITICAL (fn () => (Keyword.dest_keywords (), #2 (get_syntax ()))); | |
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 171 | fun pretty_cmd (name, comment, _) = | 
| 7026 | 172 | Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; | 
| 43711 | 173 | val (int_cmds, cmds) = List.partition #3 (dest_commands outer_syntax); | 
| 7026 | 174 | in | 
| 43711 | 175 |     [Pretty.strs ("syntax keywords:" :: map quote keywords),
 | 
| 18326 | 176 | Pretty.big_list "commands:" (map pretty_cmd cmds), | 
| 177 | Pretty.big_list "interactive-only commands:" (map pretty_cmd int_cmds)] | |
| 9223 | 178 | |> Pretty.chunks |> Pretty.writeln | 
| 7026 | 179 | end; | 
| 5829 | 180 | |
| 181 | ||
| 182 | ||
| 9132 | 183 | (** toplevel parsing **) | 
| 5829 | 184 | |
| 9132 | 185 | (* basic sources *) | 
| 6860 | 186 | |
| 26620 | 187 | fun toplevel_source term do_recover cmd src = | 
| 9132 | 188 | let | 
| 189 | val no_terminator = | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36955diff
changeset | 190 | 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: 
23679diff
changeset | 191 | fun recover int = | 
| 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 wenzelm parents: 
23679diff
changeset | 192 | (int, fn _ => Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]); | 
| 9132 | 193 | in | 
| 194 | src | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36955diff
changeset | 195 | |> Token.source_proper | 
| 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36955diff
changeset | 196 | |> Source.source Token.stopper | 
| 36950 | 197 | (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: 
23679diff
changeset | 198 | (Option.map recover do_recover) | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19060diff
changeset | 199 | |> 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: 
36955diff
changeset | 200 | |> Source.source Token.stopper | 
| 36950 | 201 | (Scan.bulk (fn xs => Parse.!!! (parse_command term (cmd ())) xs)) | 
| 23682 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 wenzelm parents: 
23679diff
changeset | 202 | (Option.map recover do_recover) | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19060diff
changeset | 203 | |> Source.map_filter I | 
| 9132 | 204 | end; | 
| 5829 | 205 | |
| 7746 | 206 | |
| 25580 | 207 | (* off-line scanning/parsing *) | 
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 208 | |
| 27839 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 209 | fun scan pos str = | 
| 16195 | 210 | Source.of_string str | 
| 40523 
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
 wenzelm parents: 
38422diff
changeset | 211 | |> Symbol.source | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36955diff
changeset | 212 |   |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
 | 
| 16195 | 213 | |> Source.exhaust; | 
| 214 | ||
| 25580 | 215 | fun parse pos str = | 
| 216 | Source.of_string str | |
| 40523 
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
 wenzelm parents: 
38422diff
changeset | 217 | |> Symbol.source | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36955diff
changeset | 218 |   |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
 | 
| 43711 | 219 | |> toplevel_source false NONE lookup_commands_dynamic | 
| 25580 | 220 | |> Source.exhaust; | 
| 221 | ||
| 14091 | 222 | |
| 26431 | 223 | (* process file *) | 
| 224 | ||
| 225 | fun process_file path thy = | |
| 226 | let | |
| 26881 | 227 | val trs = parse (Path.position path) (File.read path); | 
| 44187 
88d770052bac
simplified Toplevel.init_theory: discontinued special name argument;
 wenzelm parents: 
44186diff
changeset | 228 | val init = Toplevel.init_theory (K thy) Toplevel.empty; | 
| 28424 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 wenzelm parents: 
27872diff
changeset | 229 | val result = fold Toplevel.command (init :: trs) Toplevel.toplevel; | 
| 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 wenzelm parents: 
27872diff
changeset | 230 | in | 
| 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 wenzelm parents: 
27872diff
changeset | 231 | (case (Toplevel.is_theory result, Toplevel.generic_theory_of result) of | 
| 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 wenzelm parents: 
27872diff
changeset | 232 | (true, Context.Theory thy') => thy' | 
| 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 wenzelm parents: 
27872diff
changeset | 233 | | _ => error "Bad result state: global theory expected") | 
| 
fc6ce1c4d5b7
simplified process_file, eliminated Toplevel.excursion;
 wenzelm parents: 
27872diff
changeset | 234 | end; | 
| 26431 | 235 | |
| 236 | ||
| 24868 | 237 | (* interactive source of toplevel transformers *) | 
| 238 | ||
| 26600 | 239 | type isar = | 
| 240 | (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: 
36955diff
changeset | 241 | (Token.T, (Token.T option, (Token.T, (Token.T, | 
| 30573 | 242 | (Symbol_Pos.T, Position.T * (Symbol.symbol, (string, unit) Source.source) | 
| 27770 | 243 | Source.source) Source.source) Source.source) Source.source) | 
| 244 | Source.source) Source.source) Source.source) Source.source; | |
| 26600 | 245 | |
| 38253 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
38236diff
changeset | 246 | fun isar in_stream term : isar = | 
| 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
38236diff
changeset | 247 | Source.tty in_stream | 
| 40523 
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
 wenzelm parents: 
38422diff
changeset | 248 | |> Symbol.source | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36955diff
changeset | 249 |   |> Token.source {do_recover = SOME true} Keyword.get_lexicons Position.none
 | 
| 43711 | 250 | |> toplevel_source term (SOME true) lookup_commands_dynamic; | 
| 24868 | 251 | |
| 252 | ||
| 44478 
4fdb1009a370
tuned signature -- emphasize traditional read/eval/print terminology, which is still relevant here;
 wenzelm parents: 
44357diff
changeset | 253 | (* read toplevel commands -- fail-safe *) | 
| 27839 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 254 | |
| 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 255 | val not_singleton = "Exactly one command expected"; | 
| 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 256 | |
| 44658 
5bec9c15ef29
more direct Token.range_pos and Outer_Syntax.read_command, bypassing Thy_Syntax.span;
 wenzelm parents: 
44478diff
changeset | 257 | fun read_span outer_syntax toks = | 
| 27839 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 258 | let | 
| 43711 | 259 | val commands = lookup_commands outer_syntax; | 
| 44658 
5bec9c15ef29
more direct Token.range_pos and Outer_Syntax.read_command, bypassing Thy_Syntax.span;
 wenzelm parents: 
44478diff
changeset | 260 | val range_pos = Position.set_range (Token.range toks); | 
| 44736 | 261 | val _ = Position.reports (maps Thy_Syntax.reports_of_token toks); | 
| 27839 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 262 | in | 
| 29311 | 263 | (case Source.exhaust (toplevel_source false NONE (K commands) (Source.of_list toks)) of | 
| 37713 
c82cf6e11669
Outer_Syntax.prepare_command: disallow control commands here, and consequently in Isar documents;
 wenzelm parents: 
37216diff
changeset | 264 | [tr] => | 
| 
c82cf6e11669
Outer_Syntax.prepare_command: disallow control commands here, and consequently in Isar documents;
 wenzelm parents: 
37216diff
changeset | 265 | if Keyword.is_control (Toplevel.name_of tr) then | 
| 44658 
5bec9c15ef29
more direct Token.range_pos and Outer_Syntax.read_command, bypassing Thy_Syntax.span;
 wenzelm parents: 
44478diff
changeset | 266 | (Toplevel.malformed (Toplevel.pos_of tr) "Illegal control command", true) | 
| 37713 
c82cf6e11669
Outer_Syntax.prepare_command: disallow control commands here, and consequently in Isar documents;
 wenzelm parents: 
37216diff
changeset | 267 | else (tr, true) | 
| 27839 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 268 | | [] => (Toplevel.ignored range_pos, false) | 
| 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 269 | | _ => (Toplevel.malformed range_pos not_singleton, true)) | 
| 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 270 | handle ERROR msg => (Toplevel.malformed range_pos msg, true) | 
| 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 271 | end; | 
| 
0be8644c45eb
Symbol.source/OuterLex.source: more explicit do_recover argument;
 wenzelm parents: 
27831diff
changeset | 272 | |
| 44659 | 273 | fun read_command pos str = | 
| 274 | let | |
| 275 | val (lexs, outer_syntax) = get_syntax (); | |
| 276 | val toks = Thy_Syntax.parse_tokens lexs pos str; | |
| 277 | in #1 (read_span outer_syntax toks) end; | |
| 278 | ||
| 44478 
4fdb1009a370
tuned signature -- emphasize traditional read/eval/print terminology, which is still relevant here;
 wenzelm parents: 
44357diff
changeset | 279 | fun read_element outer_syntax init {head, proof, proper_proof} =
 | 
| 28436 
4faf705a177d
load_thy: more precise treatment of improper cmd or proof (notably 'oops');
 wenzelm parents: 
28432diff
changeset | 280 | let | 
| 44658 
5bec9c15ef29
more direct Token.range_pos and Outer_Syntax.read_command, bypassing Thy_Syntax.span;
 wenzelm parents: 
44478diff
changeset | 281 | val read = read_span outer_syntax o Thy_Syntax.span_content; | 
| 
5bec9c15ef29
more direct Token.range_pos and Outer_Syntax.read_command, bypassing Thy_Syntax.span;
 wenzelm parents: 
44478diff
changeset | 282 | val (tr, proper_head) = read head |>> Toplevel.modify_init init; | 
| 
5bec9c15ef29
more direct Token.range_pos and Outer_Syntax.read_command, bypassing Thy_Syntax.span;
 wenzelm parents: 
44478diff
changeset | 283 | val proof_trs = map read proof |> filter #2 |> map #1; | 
| 28436 
4faf705a177d
load_thy: more precise treatment of improper cmd or proof (notably 'oops');
 wenzelm parents: 
28432diff
changeset | 284 | in | 
| 43621 | 285 | if proper_head andalso proper_proof then [(tr, proof_trs)] | 
| 286 | else map (rpair []) (if proper_head then tr :: proof_trs else proof_trs) | |
| 28436 
4faf705a177d
load_thy: more precise treatment of improper cmd or proof (notably 'oops');
 wenzelm parents: 
28432diff
changeset | 287 | end; | 
| 28432 
944cb67f809a
load_thy: Toplevel.excursion based on units of command/proof pairs;
 wenzelm parents: 
28424diff
changeset | 288 | |
| 5829 | 289 | end; | 
| 36953 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 wenzelm parents: 
36950diff
changeset | 290 |