| author | wenzelm | 
| Mon, 20 Jun 2005 22:14:02 +0200 | |
| changeset 16488 | 38bc902946b2 | 
| parent 16195 | 0eb3c15298cd | 
| child 16727 | e264077b68a7 | 
| permissions | -rw-r--r-- | 
| 5829 | 1 | (* Title: Pure/Isar/outer_syntax.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 5 | The global Isabelle/Isar outer syntax. | |
| 6 | *) | |
| 7 | ||
| 8 | signature BASIC_OUTER_SYNTAX = | |
| 9 | sig | |
| 15830 | 10 | structure Isar: | 
| 11 | sig | |
| 12 | val main: unit -> unit | |
| 13 | val loop: unit -> unit | |
| 14 | val sync_main: unit -> unit | |
| 15 | val sync_loop: unit -> unit | |
| 16 | end; | |
| 5829 | 17 | end; | 
| 18 | ||
| 19 | signature OUTER_SYNTAX = | |
| 20 | sig | |
| 21 | include BASIC_OUTER_SYNTAX | |
| 6722 | 22 | structure Keyword: | 
| 23 | sig | |
| 24 | val control: string | |
| 25 | val diag: string | |
| 26 | val thy_begin: string | |
| 7104 | 27 | val thy_switch: string | 
| 6722 | 28 | val thy_end: string | 
| 29 | val thy_heading: string | |
| 30 | val thy_decl: string | |
| 9588 | 31 | val thy_script: string | 
| 6722 | 32 | val thy_goal: string | 
| 33 | val qed: string | |
| 6733 | 34 | val qed_block: string | 
| 8209 | 35 | val qed_global: string | 
| 9552 | 36 | val prf_heading: string | 
| 6722 | 37 | val prf_goal: string | 
| 38 | val prf_block: string | |
| 9056 | 39 | val prf_open: string | 
| 40 | val prf_close: string | |
| 6722 | 41 | val prf_chain: string | 
| 42 | val prf_decl: string | |
| 6868 | 43 | val prf_asm: string | 
| 7676 | 44 | val prf_asm_goal: string | 
| 6722 | 45 | val prf_script: string | 
| 46 | val kinds: string list | |
| 47 | end | |
| 5829 | 48 | type token | 
| 49 | type parser | |
| 6722 | 50 | val command: string -> string -> string -> | 
| 6373 | 51 | (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser | 
| 9132 | 52 | val markup_command: IsarOutput.markup -> string -> string -> string -> | 
| 7789 | 53 | (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser | 
| 6722 | 54 | val improper_command: string -> string -> string -> | 
| 6373 | 55 | (token list -> (Toplevel.transition -> Toplevel.transition) * token list) -> parser | 
| 14687 | 56 | val is_keyword: string -> bool | 
| 7026 | 57 | val dest_keywords: unit -> string list | 
| 58 | val dest_parsers: unit -> (string * string * string * bool) list | |
| 5883 | 59 | val print_outer_syntax: unit -> unit | 
| 9223 | 60 | val print_commands: Toplevel.transition -> Toplevel.transition | 
| 5829 | 61 | val add_keywords: string list -> unit | 
| 62 | val add_parsers: parser list -> unit | |
| 12943 | 63 | val check_text: string * Position.T -> bool -> Toplevel.state -> unit | 
| 7940 | 64 | val deps_thy: string -> bool -> Path.T -> string list * Path.T list | 
| 6199 | 65 | val load_thy: string -> bool -> bool -> Path.T -> unit | 
| 14091 | 66 | val isar: bool -> bool -> unit Toplevel.isar | 
| 16195 | 67 | val scan: string -> OuterLex.token list | 
| 68 | val read: OuterLex.token list -> (string * OuterLex.token list * Toplevel.transition) list | |
| 5829 | 69 | end; | 
| 70 | ||
| 15224 
1bd35fd87963
Allow scanning to recover and reconstruct bad input
 aspinall parents: 
15156diff
changeset | 71 | structure OuterSyntax : OUTER_SYNTAX = | 
| 5829 | 72 | struct | 
| 73 | ||
| 7750 | 74 | structure T = OuterLex; | 
| 6860 | 75 | structure P = OuterParse; | 
| 76 | ||
| 5829 | 77 | |
| 78 | (** outer syntax **) | |
| 79 | ||
| 6722 | 80 | (* command keyword classification *) | 
| 81 | ||
| 82 | structure Keyword = | |
| 83 | struct | |
| 84 | val control = "control"; | |
| 85 | val diag = "diag"; | |
| 86 | val thy_begin = "theory-begin"; | |
| 7104 | 87 | val thy_switch = "theory-switch"; | 
| 6722 | 88 | val thy_end = "theory-end"; | 
| 89 | val thy_heading = "theory-heading"; | |
| 90 | val thy_decl = "theory-decl"; | |
| 9588 | 91 | val thy_script = "theory-script"; | 
| 6722 | 92 | val thy_goal = "theory-goal"; | 
| 93 | val qed = "qed"; | |
| 6733 | 94 | val qed_block = "qed-block"; | 
| 8209 | 95 | val qed_global = "qed-global"; | 
| 9552 | 96 | val prf_heading = "proof-heading"; | 
| 6722 | 97 | val prf_goal = "proof-goal"; | 
| 98 | val prf_block = "proof-block"; | |
| 9056 | 99 | val prf_open = "proof-open"; | 
| 100 | val prf_close = "proof-close"; | |
| 6722 | 101 | val prf_chain = "proof-chain"; | 
| 102 | val prf_decl = "proof-decl"; | |
| 6868 | 103 | val prf_asm = "proof-asm"; | 
| 7676 | 104 | val prf_asm_goal = "proof-asm-goal"; | 
| 6722 | 105 | val prf_script = "proof-script"; | 
| 106 | ||
| 9588 | 107 | val kinds = [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, | 
| 108 | thy_goal, qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open, prf_close, | |
| 109 | prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script]; | |
| 6722 | 110 | end; | 
| 111 | ||
| 112 | ||
| 5829 | 113 | (* parsers *) | 
| 114 | ||
| 7750 | 115 | type token = T.token; | 
| 5829 | 116 | type parser_fn = token list -> (Toplevel.transition -> Toplevel.transition) * token list; | 
| 117 | ||
| 118 | datatype parser = | |
| 9132 | 119 | Parser of string * (string * string * IsarOutput.markup option) * bool * parser_fn; | 
| 5829 | 120 | |
| 7750 | 121 | fun parser int_only markup name comment kind parse = | 
| 122 | Parser (name, (comment, kind, markup), int_only, parse); | |
| 5829 | 123 | |
| 124 | ||
| 125 | (* parse command *) | |
| 126 | ||
| 6860 | 127 | local | 
| 6199 | 128 | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 129 | fun terminate false = Scan.succeed () | 
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 130 | | 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 | 131 | |
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 132 | fun trace false parse = parse | 
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 133 | | 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 | 134 | |
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 135 | fun body cmd trc (name, _) = | 
| 7026 | 136 | (case cmd name of | 
| 15531 | 137 | SOME (int_only, parse) => | 
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 138 | P.!!! (Scan.prompt (name ^ "# ") (trace trc parse >> pair int_only)) | 
| 15531 | 139 |   | NONE => sys_error ("no parser for outer syntax command " ^ quote name));
 | 
| 6860 | 140 | |
| 141 | in | |
| 5829 | 142 | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 143 | fun command do_terminate do_trace cmd = | 
| 15531 | 144 | P.semicolon >> K NONE || | 
| 145 | P.sync >> K NONE || | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 146 | (P.position P.command :-- body cmd do_trace) --| terminate do_terminate | 
| 6860 | 147 | >> (fn ((name, pos), (int_only, f)) => | 
| 15531 | 148 | SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |> | 
| 6860 | 149 | Toplevel.interactive int_only |> f)); | 
| 5829 | 150 | |
| 6199 | 151 | end; | 
| 152 | ||
| 5829 | 153 | |
| 154 | ||
| 9132 | 155 | (** global outer syntax **) | 
| 5829 | 156 | |
| 7026 | 157 | local | 
| 158 | ||
| 159 | val global_lexicons = ref (Scan.empty_lexicon, Scan.empty_lexicon); | |
| 7750 | 160 | val global_parsers = | 
| 9132 | 161 | ref (Symtab.empty: (((string * string) * (bool * parser_fn)) * IsarOutput.markup option) | 
| 162 | Symtab.table); | |
| 163 | val global_markups = ref ([]: (string * IsarOutput.markup) list); | |
| 5952 | 164 | |
| 7026 | 165 | fun change_lexicons f = | 
| 166 | let val lexs = f (! global_lexicons) in | |
| 167 | (case (op inter_string) (pairself Scan.dest_lexicon lexs) of | |
| 168 | [] => global_lexicons := lexs | |
| 169 |     | bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads))
 | |
| 170 | end; | |
| 5829 | 171 | |
| 15531 | 172 | fun get_markup (ms, (name, (_, SOME m))) = (name, m) :: ms | 
| 7789 | 173 | | get_markup (ms, _) = ms; | 
| 7750 | 174 | |
| 175 | fun make_markups () = global_markups := Symtab.foldl get_markup ([], ! global_parsers); | |
| 9132 | 176 | fun change_parsers f = (Library.change global_parsers f; make_markups ()); | 
| 6722 | 177 | |
| 7026 | 178 | in | 
| 179 | ||
| 7750 | 180 | |
| 9132 | 181 | (* access current syntax *) | 
| 7026 | 182 | |
| 183 | (*Note: the syntax for files is statically determined at the very | |
| 184 | beginning; for interactive processing it may change dynamically.*) | |
| 185 | ||
| 186 | fun get_lexicons () = ! global_lexicons; | |
| 187 | fun get_parsers () = ! global_parsers; | |
| 15570 | 188 | fun get_parser () = Option.map (#2 o #1) o curry Symtab.lookup (! global_parsers); | 
| 7789 | 189 | |
| 9132 | 190 | fun is_markup kind name = | 
| 15531 | 191 | (case assoc (! global_markups, name) of SOME k => k = kind | NONE => false); | 
| 9132 | 192 | fun markup kind = Scan.one (T.is_kind T.Command andf is_markup kind o T.val_of); | 
| 5829 | 193 | |
| 194 | ||
| 195 | (* augment syntax *) | |
| 196 | ||
| 7026 | 197 | fun add_keywords keywords = change_lexicons (apfst (fn lex => | 
| 198 | (Scan.extend_lexicon lex (map Symbol.explode keywords)))); | |
| 5829 | 199 | |
| 7750 | 200 | fun add_parser (tab, Parser (name, (comment, kind, markup), int_only, parse)) = | 
| 5829 | 201 | (if is_none (Symtab.lookup (tab, name)) then () | 
| 202 |   else warning ("Redefined outer syntax command " ^ quote name);
 | |
| 7750 | 203 | Symtab.update ((name, (((comment, kind), (int_only, parse)), markup)), tab)); | 
| 5829 | 204 | |
| 205 | fun add_parsers parsers = | |
| 15570 | 206 | (change_parsers (fn tab => Library.foldl add_parser (tab, parsers)); | 
| 7026 | 207 | change_lexicons (apsnd (fn lex => Scan.extend_lexicon lex | 
| 208 | (map (fn Parser (name, _, _, _) => Symbol.explode name) parsers)))); | |
| 209 | ||
| 210 | end; | |
| 5829 | 211 | |
| 212 | ||
| 7026 | 213 | (* print syntax *) | 
| 214 | ||
| 14687 | 215 | fun is_keyword s = Scan.is_literal (#1 (get_lexicons ())) (Symbol.explode s); | 
| 7026 | 216 | fun dest_keywords () = Scan.dest_lexicon (#1 (get_lexicons ())); | 
| 217 | ||
| 218 | fun dest_parsers () = | |
| 7750 | 219 | map (fn (name, (((cmt, kind), (int_only, _)), _)) => (name, cmt, kind, int_only)) | 
| 7026 | 220 | (Symtab.dest (get_parsers ())); | 
| 5829 | 221 | |
| 9223 | 222 | fun print_outer_syntax () = | 
| 7026 | 223 | let | 
| 224 | fun pretty_cmd (name, comment, _, _) = | |
| 225 | Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; | |
| 15570 | 226 | val (int_cmds, cmds) = List.partition #4 (dest_parsers ()); | 
| 7026 | 227 | in | 
| 8720 | 228 |     [Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())),
 | 
| 229 | Pretty.big_list "proper commands:" (map pretty_cmd cmds), | |
| 230 | Pretty.big_list "improper commands (interactive-only):" (map pretty_cmd int_cmds)] | |
| 9223 | 231 | |> Pretty.chunks |> Pretty.writeln | 
| 7026 | 232 | end; | 
| 5829 | 233 | |
| 9223 | 234 | val print_commands = Toplevel.imperative print_outer_syntax; | 
| 7367 | 235 | |
| 5829 | 236 | |
| 237 | ||
| 9132 | 238 | (** toplevel parsing **) | 
| 5829 | 239 | |
| 9132 | 240 | (* basic sources *) | 
| 6860 | 241 | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 242 | fun toplevel_source term trc do_recover cmd src = | 
| 9132 | 243 | let | 
| 244 | val no_terminator = | |
| 245 | Scan.unless P.semicolon (Scan.one (T.not_sync andf T.not_eof)); | |
| 15531 | 246 | fun recover x = (Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]) x; | 
| 9132 | 247 | in | 
| 248 | src | |
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
10749diff
changeset | 249 | |> T.source_proper | 
| 9132 | 250 | |> Source.source T.stopper | 
| 15531 | 251 | (Scan.bulk (P.$$$ "--" -- P.!!! P.text >> K NONE || P.not_eof >> SOME)) | 
| 252 | (if do_recover then SOME recover else NONE) | |
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
10749diff
changeset | 253 | |> Source.mapfilter I | 
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 254 | |> Source.source T.stopper (Scan.bulk (fn xs => P.!!! (command term trc (cmd ())) xs)) | 
| 15531 | 255 | (if do_recover then SOME recover else NONE) | 
| 9132 | 256 | |> Source.mapfilter I | 
| 257 | end; | |
| 5829 | 258 | |
| 7746 | 259 | |
| 9132 | 260 | (* interactive source of toplevel transformers *) | 
| 5829 | 261 | |
| 9132 | 262 | fun isar term no_pos = | 
| 263 | Source.tty | |
| 264 | |> Symbol.source true | |
| 265 | |> T.source true get_lexicons | |
| 266 | (if no_pos then Position.none else Position.line_name 1 "stdin") | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 267 | |> toplevel_source term false true get_parser; | 
| 6199 | 268 | |
| 269 | ||
| 16195 | 270 | (* scan text *) | 
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 271 | |
| 15144 
85929e1b307d
Remove isar_readstring.  Split read into scanner and parser.
 aspinall parents: 
14981diff
changeset | 272 | fun scan str = | 
| 16195 | 273 | Source.of_string str | 
| 274 | |> Symbol.source false | |
| 275 | |> T.source true get_lexicons Position.none | |
| 276 | |> Source.exhaust; | |
| 277 | ||
| 278 | ||
| 279 | (* read tokens with trace *) | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 280 | |
| 15144 
85929e1b307d
Remove isar_readstring.  Split read into scanner and parser.
 aspinall parents: 
14981diff
changeset | 281 | fun read toks = | 
| 
85929e1b307d
Remove isar_readstring.  Split read into scanner and parser.
 aspinall parents: 
14981diff
changeset | 282 | Source.of_list toks | 
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 283 | |> toplevel_source false true true get_parser | 
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 284 | |> Source.exhaust | 
| 15973 | 285 | |> map (fn tr => (Toplevel.name_of tr, the (Toplevel.source_of tr), tr)); | 
| 14091 | 286 | |
| 287 | ||
| 16195 | 288 | |
| 9132 | 289 | (** read theory **) | 
| 6247 | 290 | |
| 12943 | 291 | (* check_text *) | 
| 292 | ||
| 293 | fun check_text s true state = (IsarOutput.eval_antiquote (#1 (get_lexicons ())) state s; ()) | |
| 294 | | check_text _ false _ = (); | |
| 295 | ||
| 296 | ||
| 9132 | 297 | (* deps_thy *) | 
| 6199 | 298 | |
| 7940 | 299 | fun deps_thy name ml path = | 
| 6199 | 300 | let | 
| 7735 | 301 | val src = Source.of_string (File.read path); | 
| 302 | val pos = Path.position path; | |
| 9132 | 303 | val (name', parents, files) = ThyHeader.scan (src, pos); | 
| 6199 | 304 | val ml_path = ThyLoad.ml_path name; | 
| 15973 | 305 | val ml_file = if ml andalso is_some (ThyLoad.check_file NONE ml_path) then [ml_path] else []; | 
| 6199 | 306 | in | 
| 307 | if name <> name' then | |
| 7940 | 308 |       error ("Filename " ^ quote (Path.pack path) ^
 | 
| 8078 | 309 | " does not agree with theory name " ^ quote name') | 
| 6247 | 310 | else (parents, map (Path.unpack o #1) files @ ml_file) | 
| 6199 | 311 | end; | 
| 312 | ||
| 7746 | 313 | |
| 9132 | 314 | (* load_thy *) | 
| 6199 | 315 | |
| 7746 | 316 | local | 
| 317 | ||
| 7940 | 318 | fun try_ml_file name time = | 
| 6199 | 319 | let | 
| 320 | val path = ThyLoad.ml_path name; | |
| 7940 | 321 | val tr = Toplevel.imperative (fn () => ThyInfo.load_file time path); | 
| 6247 | 322 | val tr_name = if time then "time_use" else "use"; | 
| 6199 | 323 | in | 
| 15531 | 324 | if is_none (ThyLoad.check_file NONE path) then () | 
| 9132 | 325 | else Toplevel.excursion [Toplevel.empty |> Toplevel.name tr_name |> tr] | 
| 6199 | 326 | end; | 
| 327 | ||
| 9132 | 328 | fun parse_thy src = | 
| 329 | src | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 330 | |> toplevel_source false false false (K (get_parser ())) | 
| 7746 | 331 | |> Source.exhaust; | 
| 5829 | 332 | |
| 6247 | 333 | fun run_thy name path = | 
| 7683 | 334 | let | 
| 7735 | 335 | val pos = Path.position path; | 
| 9132 | 336 | val text = Library.untabify (explode (File.read path)); | 
| 337 | val text_src = Source.of_list text; | |
| 338 | fun present_text () = Source.exhaust (Symbol.source false text_src); | |
| 7683 | 339 | in | 
| 7735 | 340 | Present.init_theory name; | 
| 9132 | 341 | Present.verbatim_source name present_text; | 
| 342 | if ThyHeader.is_old (text_src, pos) then (ThySyn.load_thy name text; | |
| 343 | Present.old_symbol_source name present_text) (*note: text presented twice*) | |
| 344 | else | |
| 345 | let | |
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
10749diff
changeset | 346 | val tok_src = text_src | 
| 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
10749diff
changeset | 347 | |> Symbol.source false | 
| 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
10749diff
changeset | 348 | |> T.source false (K (get_lexicons ())) pos | 
| 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
10749diff
changeset | 349 | |> Source.exhausted; | 
| 9132 | 350 | val out = Toplevel.excursion_result | 
| 351 | (IsarOutput.parse_thy markup (#1 (get_lexicons ())) | |
| 352 | (parse_thy tok_src) tok_src); | |
| 353 | in Present.theory_output name (Buffer.content out) end | |
| 6247 | 354 | end; | 
| 6199 | 355 | |
| 7746 | 356 | in | 
| 357 | ||
| 6199 | 358 | fun load_thy name ml time path = | 
| 6247 | 359 | (if time then | 
| 360 | timeit (fn () => | |
| 361 |      (writeln ("\n**** Starting theory " ^ quote name ^ " ****");
 | |
| 9036 | 362 | run_thy name path; | 
| 6247 | 363 |       writeln ("**** Finished theory " ^ quote name ^ " ****\n")))
 | 
| 364 | else run_thy name path; | |
| 365 | Context.context (ThyInfo.get_theory name); | |
| 7940 | 366 | if ml then try_ml_file name time else ()); | 
| 5829 | 367 | |
| 7746 | 368 | end; | 
| 369 | ||
| 5829 | 370 | |
| 371 | ||
| 372 | (** the read-eval-print loop **) | |
| 373 | ||
| 5923 | 374 | (* main loop *) | 
| 375 | ||
| 7333 | 376 | fun gen_loop term no_pos = | 
| 377 | (Context.reset_context (); | |
| 15989 | 378 | Toplevel.loop (isar term no_pos); | 
| 379 | ml_prompts "ML> " "ML# "); | |
| 5829 | 380 | |
| 7333 | 381 | fun gen_main term no_pos = | 
| 5829 | 382 | (Toplevel.set_state Toplevel.toplevel; | 
| 6199 | 383 | writeln (Session.welcome ()); | 
| 7333 | 384 | gen_loop term no_pos); | 
| 6860 | 385 | |
| 15830 | 386 | structure Isar = | 
| 387 | struct | |
| 388 | fun main () = gen_main false false; | |
| 389 | fun loop () = gen_loop false false; | |
| 390 | fun sync_main () = gen_main true true; | |
| 391 | fun sync_loop () = gen_loop true true; | |
| 392 | end; | |
| 5829 | 393 | |
| 394 | ||
| 6373 | 395 | (*final declarations of this structure!*) | 
| 15531 | 396 | val command = parser false NONE; | 
| 397 | val markup_command = parser false o SOME; | |
| 398 | val improper_command = parser true NONE; | |
| 6685 | 399 | |
| 5829 | 400 | end; | 
| 401 | ||
| 6199 | 402 | (*setup theory syntax dependent operations*) | 
| 403 | ThyLoad.deps_thy_fn := OuterSyntax.deps_thy; | |
| 404 | ThyLoad.load_thy_fn := OuterSyntax.load_thy; | |
| 405 | structure ThyLoad: THY_LOAD = ThyLoad; | |
| 406 | ||
| 5829 | 407 | structure BasicOuterSyntax: BASIC_OUTER_SYNTAX = OuterSyntax; | 
| 408 | open BasicOuterSyntax; | |
| 15830 | 409 | open Isar; |