| author | wenzelm | 
| Wed, 10 Oct 2007 17:32:00 +0200 | |
| changeset 24952 | f336c36f41a0 | 
| parent 24872 | 7fd1aa6671a4 | 
| child 25526 | 05ee7d85912e | 
| permissions | -rw-r--r-- | 
| 5829 | 1 | (* Title: Pure/Isar/outer_syntax.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 24868 | 5 | The global Isabelle/Isar outer syntax. Note: the syntax for files is | 
| 6 | statically determined at the very beginning; for interactive processing | |
| 7 | it may change dynamically. | |
| 5829 | 8 | *) | 
| 9 | ||
| 10 | signature BASIC_OUTER_SYNTAX = | |
| 11 | sig | |
| 15830 | 12 | structure Isar: | 
| 13 | sig | |
| 18064 | 14 | val state: unit -> Toplevel.state | 
| 21207 | 15 | val exn: unit -> (exn * string) option | 
| 20023 | 16 | val context: unit -> Proof.context | 
| 21401 | 17 | val goal: unit -> thm list * thm | 
| 15830 | 18 | val main: unit -> unit | 
| 19 | val loop: unit -> unit | |
| 20 | val sync_main: unit -> unit | |
| 21 | val sync_loop: unit -> unit | |
| 18684 | 22 | val toplevel: (unit -> 'a) -> 'a | 
| 15830 | 23 | end; | 
| 5829 | 24 | end; | 
| 25 | ||
| 26 | signature OUTER_SYNTAX = | |
| 27 | sig | |
| 28 | include BASIC_OUTER_SYNTAX | |
| 24868 | 29 | type parser_fn = OuterLex.token list -> | 
| 30 | (Toplevel.transition -> Toplevel.transition) * OuterLex.token list | |
| 23722 | 31 | val get_lexicons: unit -> Scan.lexicon * Scan.lexicon | 
| 23796 | 32 | val command_keyword: string -> OuterKeyword.T option | 
| 14687 | 33 | val is_keyword: string -> bool | 
| 24868 | 34 | val keywords: string list -> unit | 
| 35 | val command: string -> string -> OuterKeyword.T -> parser_fn -> unit | |
| 36 | val markup_command: ThyOutput.markup -> string -> string -> OuterKeyword.T -> parser_fn -> unit | |
| 37 | val improper_command: string -> string -> OuterKeyword.T -> parser_fn -> unit | |
| 7026 | 38 | val dest_keywords: unit -> string list | 
| 39 | val dest_parsers: unit -> (string * string * string * bool) list | |
| 5883 | 40 | val print_outer_syntax: unit -> unit | 
| 24872 | 41 | val report: unit -> unit | 
| 19060 | 42 | val check_text: string * Position.T -> Toplevel.node option -> unit | 
| 16195 | 43 | val scan: string -> OuterLex.token list | 
| 44 | val read: OuterLex.token list -> (string * OuterLex.token list * Toplevel.transition) list | |
| 24868 | 45 | val isar: bool -> unit Toplevel.isar | 
| 5829 | 46 | end; | 
| 47 | ||
| 15224 
1bd35fd87963
Allow scanning to recover and reconstruct bad input
 aspinall parents: 
15156diff
changeset | 48 | structure OuterSyntax : OUTER_SYNTAX = | 
| 5829 | 49 | struct | 
| 50 | ||
| 7750 | 51 | structure T = OuterLex; | 
| 6860 | 52 | structure P = OuterParse; | 
| 53 | ||
| 5829 | 54 | |
| 55 | (** outer syntax **) | |
| 56 | ||
| 24872 | 57 | (* diagnostics *) | 
| 58 | ||
| 59 | fun report_keyword name = | |
| 60 | Pretty.markup (Markup.keyword_decl name) | |
| 61 |     [Pretty.str ("Outer syntax keyword: " ^ quote name)];
 | |
| 62 | ||
| 63 | fun report_command name kind = | |
| 64 | Pretty.markup (Markup.command_decl name kind) | |
| 65 |     [Pretty.str ("Outer syntax command: " ^ quote name ^ " (" ^ kind ^ ")")];
 | |
| 66 | ||
| 67 | ||
| 5829 | 68 | (* parsers *) | 
| 69 | ||
| 24868 | 70 | type parser_fn = T.token list -> (Toplevel.transition -> Toplevel.transition) * T.token list; | 
| 5829 | 71 | |
| 24868 | 72 | datatype parser = Parser of | 
| 73 |  {comment: string,
 | |
| 74 | kind: OuterKeyword.T, | |
| 75 | markup: ThyOutput.markup option, | |
| 76 | int_only: bool, | |
| 77 | parse: parser_fn}; | |
| 5829 | 78 | |
| 24868 | 79 | fun make_parser comment kind markup int_only parse = | 
| 80 |   Parser {comment = comment, kind = kind, markup = markup, int_only = int_only, parse = parse};
 | |
| 5829 | 81 | |
| 82 | ||
| 83 | (* parse command *) | |
| 84 | ||
| 6860 | 85 | local | 
| 6199 | 86 | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 87 | fun terminate false = Scan.succeed () | 
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 88 | | terminate true = P.group "end of input" (Scan.option P.sync -- P.semicolon >> K ()); | 
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 89 | |
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 90 | fun trace false parse = parse | 
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 91 | | trace true parse = Scan.trace parse >> (fn (f, toks) => f o Toplevel.source toks); | 
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 92 | |
| 17071 
f753d6dd9bd0
moved structure Keyword to OuterKeyword (Isar/outer_keyword.ML);
 wenzelm parents: 
16894diff
changeset | 93 | fun body cmd do_trace (name, _) = | 
| 7026 | 94 | (case cmd name of | 
| 24868 | 95 |     SOME (Parser {int_only, parse, ...}) =>
 | 
| 17118 | 96 | P.!!! (Scan.prompt (name ^ "# ") (trace do_trace (P.tags |-- parse) >> pair int_only)) | 
| 15531 | 97 |   | NONE => sys_error ("no parser for outer syntax command " ^ quote name));
 | 
| 6860 | 98 | |
| 99 | in | |
| 5829 | 100 | |
| 24868 | 101 | fun parse_command do_terminate do_trace cmd = | 
| 15531 | 102 | P.semicolon >> K NONE || | 
| 103 | P.sync >> K NONE || | |
| 17118 | 104 | (P.position P.command :-- body cmd do_trace) --| terminate do_terminate | 
| 6860 | 105 | >> (fn ((name, pos), (int_only, f)) => | 
| 15531 | 106 | SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |> | 
| 6860 | 107 | Toplevel.interactive int_only |> f)); | 
| 5829 | 108 | |
| 6199 | 109 | end; | 
| 110 | ||
| 5829 | 111 | |
| 112 | ||
| 9132 | 113 | (** global outer syntax **) | 
| 5829 | 114 | |
| 7026 | 115 | local | 
| 116 | ||
| 117 | val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon); | |
| 24868 | 118 | val global_parsers = ref (Symtab.empty: parser Symtab.table); | 
| 22120 | 119 | val global_markups = ref ([]: (string * ThyOutput.markup) list); | 
| 5952 | 120 | |
| 23939 | 121 | fun change_lexicons f = CRITICAL (fn () => | 
| 7026 | 122 | let val lexs = f (! global_lexicons) in | 
| 123 | (case (op inter_string) (pairself Scan.dest_lexicon lexs) of | |
| 124 | [] => global_lexicons := lexs | |
| 125 |     | bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads))
 | |
| 23939 | 126 | end); | 
| 5829 | 127 | |
| 23939 | 128 | fun change_parsers f = CRITICAL (fn () => | 
| 129 | (change global_parsers f; | |
| 130 | global_markups := | |
| 24868 | 131 |     Symtab.fold (fn (name, Parser {markup = SOME m, ...}) => cons (name, m) | _ => I)
 | 
| 132 | (! global_parsers) [])); | |
| 6722 | 133 | |
| 7026 | 134 | in | 
| 135 | ||
| 9132 | 136 | (* access current syntax *) | 
| 7026 | 137 | |
| 24868 | 138 | fun get_lexicons () = CRITICAL (fn () => ! global_lexicons); | 
| 139 | fun get_parsers () = CRITICAL (fn () => ! global_parsers); | |
| 140 | fun get_markups () = CRITICAL (fn () => ! global_markups); | |
| 7026 | 141 | |
| 24868 | 142 | fun get_parser () = Symtab.lookup (get_parsers ()); | 
| 7789 | 143 | |
| 23796 | 144 | fun command_keyword name = | 
| 24868 | 145 | (case Symtab.lookup (get_parsers ()) name of | 
| 146 |     SOME (Parser {kind, ...}) => SOME kind
 | |
| 147 | | NONE => NONE); | |
| 148 | ||
| 23796 | 149 | fun command_tags name = these ((Option.map OuterKeyword.tags_of) (command_keyword name)); | 
| 17071 
f753d6dd9bd0
moved structure Keyword to OuterKeyword (Isar/outer_keyword.ML);
 wenzelm parents: 
16894diff
changeset | 150 | |
| 24868 | 151 | fun is_markup kind name = AList.lookup (op =) (get_markups ()) name = SOME kind; | 
| 5829 | 152 | |
| 153 | ||
| 154 | (* augment syntax *) | |
| 155 | ||
| 24872 | 156 | fun keywords names = | 
| 157 | (change_lexicons (apfst (Scan.extend_lexicon (map Symbol.explode names))); | |
| 158 | List.app (Pretty.writeln o report_keyword) names); | |
| 24868 | 159 | |
| 160 | ||
| 24872 | 161 | fun add_parser (name, parser as Parser {kind, ...}) =
 | 
| 24868 | 162 | (if not (Symtab.defined (get_parsers ()) name) then () | 
| 163 |   else warning ("Redefining outer syntax command " ^ quote name);
 | |
| 164 | change_parsers (Symtab.update (name, parser)); | |
| 24872 | 165 | change_lexicons (apsnd (Scan.extend_lexicon [Symbol.explode name])); | 
| 166 | Pretty.writeln (report_command name (OuterKeyword.kind_of kind))); | |
| 5829 | 167 | |
| 24868 | 168 | fun command name comment kind parse = | 
| 169 | add_parser (name, make_parser comment kind NONE false parse); | |
| 5829 | 170 | |
| 24868 | 171 | fun markup_command markup name comment kind parse = | 
| 172 | add_parser (name, make_parser comment kind (SOME markup) false parse); | |
| 173 | ||
| 174 | fun improper_command name comment kind parse = | |
| 175 | add_parser (name, make_parser comment kind NONE true parse); | |
| 7026 | 176 | |
| 177 | end; | |
| 5829 | 178 | |
| 179 | ||
| 24872 | 180 | (* inspect syntax *) | 
| 7026 | 181 | |
| 14687 | 182 | fun is_keyword s = Scan.is_literal (#1 (get_lexicons ())) (Symbol.explode s); | 
| 7026 | 183 | fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ())); | 
| 184 | ||
| 185 | fun dest_parsers () = | |
| 16727 | 186 | get_parsers () |> Symtab.dest |> sort_wrt #1 | 
| 24868 | 187 |   |> map (fn (name, Parser {comment, kind, int_only, ...}) =>
 | 
| 188 | (name, comment, OuterKeyword.kind_of kind, int_only)); | |
| 5829 | 189 | |
| 9223 | 190 | fun print_outer_syntax () = | 
| 7026 | 191 | let | 
| 192 | fun pretty_cmd (name, comment, _, _) = | |
| 193 | Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; | |
| 15570 | 194 | val (int_cmds, cmds) = List.partition #4 (dest_parsers ()); | 
| 7026 | 195 | in | 
| 8720 | 196 |     [Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())),
 | 
| 18326 | 197 | Pretty.big_list "commands:" (map pretty_cmd cmds), | 
| 198 | Pretty.big_list "interactive-only commands:" (map pretty_cmd int_cmds)] | |
| 9223 | 199 | |> Pretty.chunks |> Pretty.writeln | 
| 7026 | 200 | end; | 
| 5829 | 201 | |
| 24872 | 202 | fun report () = | 
| 203 | (map report_keyword (dest_keywords ()) @ | |
| 204 | map (fn (name, _, kind, _) => report_command name kind) (dest_parsers ())) | |
| 205 | |> Pretty.chunks |> Pretty.writeln; | |
| 7367 | 206 | |
| 5829 | 207 | |
| 208 | ||
| 9132 | 209 | (** toplevel parsing **) | 
| 5829 | 210 | |
| 9132 | 211 | (* basic sources *) | 
| 6860 | 212 | |
| 17071 
f753d6dd9bd0
moved structure Keyword to OuterKeyword (Isar/outer_keyword.ML);
 wenzelm parents: 
16894diff
changeset | 213 | fun toplevel_source term do_trace do_recover cmd src = | 
| 9132 | 214 | let | 
| 215 | val no_terminator = | |
| 216 | Scan.unless P.semicolon (Scan.one (T.not_sync andf T.not_eof)); | |
| 23682 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 wenzelm parents: 
23679diff
changeset | 217 | fun recover int = | 
| 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 wenzelm parents: 
23679diff
changeset | 218 | (int, fn _ => Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]); | 
| 9132 | 219 | in | 
| 220 | src | |
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
10749diff
changeset | 221 | |> T.source_proper | 
| 9132 | 222 | |> Source.source T.stopper | 
| 15531 | 223 | (Scan.bulk (P.$$$ "--" -- P.!!! P.text >> K NONE || P.not_eof >> SOME)) | 
| 23682 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 wenzelm parents: 
23679diff
changeset | 224 | (Option.map recover do_recover) | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19060diff
changeset | 225 | |> Source.map_filter I | 
| 24868 | 226 | |> Source.source T.stopper | 
| 227 | (Scan.bulk (fn xs => P.!!! (parse_command term do_trace (cmd ())) xs)) | |
| 23682 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 wenzelm parents: 
23679diff
changeset | 228 | (Option.map recover do_recover) | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19060diff
changeset | 229 | |> Source.map_filter I | 
| 9132 | 230 | end; | 
| 5829 | 231 | |
| 7746 | 232 | |
| 16195 | 233 | (* scan text *) | 
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 234 | |
| 15144 
85929e1b307d
Remove isar_readstring.  Split read into scanner and parser.
 aspinall parents: 
14981diff
changeset | 235 | fun scan str = | 
| 16195 | 236 | Source.of_string str | 
| 237 | |> Symbol.source false | |
| 23679 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 238 | |> T.source (SOME false) get_lexicons Position.none | 
| 16195 | 239 | |> Source.exhaust; | 
| 240 | ||
| 241 | ||
| 242 | (* read tokens with trace *) | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 243 | |
| 15144 
85929e1b307d
Remove isar_readstring.  Split read into scanner and parser.
 aspinall parents: 
14981diff
changeset | 244 | fun read toks = | 
| 
85929e1b307d
Remove isar_readstring.  Split read into scanner and parser.
 aspinall parents: 
14981diff
changeset | 245 | Source.of_list toks | 
| 23679 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 246 | |> toplevel_source false true (SOME false) get_parser | 
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 247 | |> Source.exhaust | 
| 15973 | 248 | |> map (fn tr => (Toplevel.name_of tr, the (Toplevel.source_of tr), tr)); | 
| 14091 | 249 | |
| 250 | ||
| 24868 | 251 | (* interactive source of toplevel transformers *) | 
| 252 | ||
| 253 | fun isar term = | |
| 254 | Source.tty | |
| 255 | |> Symbol.source true | |
| 256 | |> T.source (SOME true) get_lexicons Position.none | |
| 257 | |> toplevel_source term false (SOME true) get_parser; | |
| 258 | ||
| 259 | ||
| 16195 | 260 | |
| 9132 | 261 | (** read theory **) | 
| 6247 | 262 | |
| 12943 | 263 | (* check_text *) | 
| 264 | ||
| 22120 | 265 | fun check_text s state = (ThyOutput.eval_antiquote (#1 (get_lexicons ())) state s; ()); | 
| 12943 | 266 | |
| 267 | ||
| 9132 | 268 | (* load_thy *) | 
| 6199 | 269 | |
| 7746 | 270 | local | 
| 271 | ||
| 23884 
1d39ec4fe73f
simplified ThyLoad interfaces: only one additional directory;
 wenzelm parents: 
23866diff
changeset | 272 | fun try_ml_file dir name time = | 
| 20323 | 273 | let val path = ThyLoad.ml_path name in | 
| 23884 
1d39ec4fe73f
simplified ThyLoad interfaces: only one additional directory;
 wenzelm parents: 
23866diff
changeset | 274 | if is_none (ThyLoad.check_file dir path) then () | 
| 20323 | 275 | else | 
| 276 | let | |
| 22826 | 277 |         val _ = legacy_feature ("loading attached ML script " ^ quote (Path.implode path));
 | 
| 20323 | 278 | val tr = Toplevel.imperative (fn () => ThyInfo.load_file time path); | 
| 279 | val tr_name = if time then "time_use" else "use"; | |
| 280 | in Toplevel.excursion [Toplevel.empty |> Toplevel.name tr_name |> tr] end | |
| 6199 | 281 | end; | 
| 282 | ||
| 24065 | 283 | fun run_thy dir name pos text time = | 
| 7683 | 284 | let | 
| 24065 | 285 | val text_src = Source.of_list (Library.untabify text); | 
| 23866 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 286 | |
| 17932 | 287 | val _ = Present.init_theory name; | 
| 24065 | 288 | val _ = Present.verbatim_source name (fn () => Source.exhaust (Symbol.source false text_src)); | 
| 289 | val toks = text_src | |
| 17932 | 290 | |> Symbol.source false | 
| 24065 | 291 | |> T.source NONE (K (get_lexicons ())) pos | 
| 17932 | 292 | |> Source.exhausted; | 
| 293 | val trs = toks | |
| 23679 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 294 | |> toplevel_source false false NONE (K (get_parser ())) | 
| 17932 | 295 | |> Source.exhaust; | 
| 23866 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 296 | |
| 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 297 |     val _ = if time then writeln ("\n**** Starting theory " ^ quote name ^ " ****") else ();
 | 
| 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 298 | val _ = cond_timeit time (fn () => | 
| 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 299 | ThyOutput.process_thy (#1 (get_lexicons ())) command_tags is_markup trs toks | 
| 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 300 | |> Buffer.content | 
| 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 301 | |> Present.theory_output name); | 
| 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 302 |     val _ = if time then writeln ("**** Finished theory " ^ quote name ^ " ****\n") else ();
 | 
| 24065 | 303 | in () end; | 
| 23866 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 304 | |
| 24188 | 305 | fun load_thy dir name pos text time = | 
| 24065 | 306 | (run_thy dir name pos text time; | 
| 24071 | 307 | CRITICAL (fn () => ML_Context.set_context (SOME (Context.Theory (ThyInfo.get_theory name)))); | 
| 24188 | 308 | try_ml_file dir name time); | 
| 6199 | 309 | |
| 7746 | 310 | in | 
| 311 | ||
| 23866 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 312 | val _ = ThyLoad.load_thy_fn := load_thy; | 
| 5829 | 313 | |
| 7746 | 314 | end; | 
| 315 | ||
| 5829 | 316 | |
| 317 | ||
| 318 | (** the read-eval-print loop **) | |
| 319 | ||
| 5923 | 320 | (* main loop *) | 
| 321 | ||
| 23679 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 322 | fun gen_loop term = | 
| 24071 | 323 | (CRITICAL (fn () => ML_Context.set_context NONE); | 
| 23679 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 324 | Toplevel.loop (isar term)); | 
| 5829 | 325 | |
| 23679 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 326 | fun gen_main term = | 
| 21957 | 327 | (Toplevel.init_state (); | 
| 6199 | 328 | writeln (Session.welcome ()); | 
| 23679 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 329 | gen_loop term); | 
| 6860 | 330 | |
| 15830 | 331 | structure Isar = | 
| 332 | struct | |
| 18064 | 333 | val state = Toplevel.get_state; | 
| 334 | val exn = Toplevel.exn; | |
| 21401 | 335 | |
| 21207 | 336 | fun context () = | 
| 21506 | 337 | Toplevel.context_of (state ()) | 
| 21207 | 338 | handle Toplevel.UNDEF => error "Unknown context"; | 
| 21401 | 339 | |
| 340 | fun goal () = | |
| 341 | #2 (Proof.get_goal (Toplevel.proof_of (state ()))) | |
| 342 | handle Toplevel.UNDEF => error "No goal present"; | |
| 343 | ||
| 23679 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 344 | fun main () = gen_main false; | 
| 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 345 | fun loop () = gen_loop false; | 
| 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 346 | fun sync_main () = gen_main true; | 
| 
57dceb84d1a0
toplevel_source: interactive flag indicates intermittent error_msg;
 wenzelm parents: 
22826diff
changeset | 347 | fun sync_loop () = gen_loop true; | 
| 18684 | 348 | val toplevel = Toplevel.program; | 
| 15830 | 349 | end; | 
| 5829 | 350 | |
| 351 | end; | |
| 352 | ||
| 6199 | 353 | structure ThyLoad: THY_LOAD = ThyLoad; | 
| 5829 | 354 | structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax; | 
| 355 | open BasicOuterSyntax; |