| author | wenzelm | 
| Tue, 15 Jul 2008 15:59:49 +0200 | |
| changeset 27608 | 8fd5662ccd97 | 
| parent 27575 | e540ad3fb50a | 
| child 27614 | f38c25d106a7 | 
| permissions | -rw-r--r-- | 
| 5829 | 1 | (* Title: Pure/Isar/outer_syntax.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 5 | The global Isabelle/Isar outer syntax. | 
| 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 6 | |
| 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 7 | 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 | 8 | beginning; for interactive processing it may change dynamically. | 
| 5829 | 9 | *) | 
| 10 | ||
| 11 | signature OUTER_SYNTAX = | |
| 12 | sig | |
| 24868 | 13 | type parser_fn = OuterLex.token list -> | 
| 14 | (Toplevel.transition -> Toplevel.transition) * OuterLex.token list | |
| 15 | val command: string -> string -> OuterKeyword.T -> parser_fn -> unit | |
| 16 | val markup_command: ThyOutput.markup -> string -> string -> OuterKeyword.T -> parser_fn -> unit | |
| 17 | val improper_command: string -> string -> OuterKeyword.T -> parser_fn -> unit | |
| 26990 | 18 | val local_theory: string -> string -> OuterKeyword.T -> | 
| 19 | (OuterParse.token list -> (local_theory -> local_theory) * OuterLex.token list) -> unit | |
| 20 | val local_theory_to_proof': string -> string -> OuterKeyword.T -> | |
| 21 | (OuterParse.token list -> (bool -> local_theory -> Proof.state) * OuterLex.token list) -> unit | |
| 22 | val local_theory_to_proof: string -> string -> OuterKeyword.T -> | |
| 23 | (OuterParse.token list -> (local_theory -> Proof.state) * OuterLex.token list) -> unit | |
| 5883 | 24 | val print_outer_syntax: unit -> unit | 
| 16195 | 25 | val scan: string -> OuterLex.token list | 
| 25580 | 26 | val parse: Position.T -> string -> Toplevel.transition list | 
| 26431 | 27 | val process_file: Path.T -> theory -> theory | 
| 26600 | 28 | type isar | 
| 29 | val isar: bool -> isar | |
| 26611 | 30 | val load_thy: Path.T -> string -> Position.T -> string list -> bool -> unit | 
| 5829 | 31 | end; | 
| 32 | ||
| 26600 | 33 | structure OuterSyntax: OUTER_SYNTAX = | 
| 5829 | 34 | struct | 
| 35 | ||
| 7750 | 36 | structure T = OuterLex; | 
| 6860 | 37 | structure P = OuterParse; | 
| 38 | ||
| 5829 | 39 | |
| 40 | (** outer syntax **) | |
| 41 | ||
| 42 | (* parsers *) | |
| 43 | ||
| 24868 | 44 | type parser_fn = T.token list -> (Toplevel.transition -> Toplevel.transition) * T.token list; | 
| 5829 | 45 | |
| 24868 | 46 | datatype parser = Parser of | 
| 47 |  {comment: string,
 | |
| 48 | markup: ThyOutput.markup option, | |
| 49 | int_only: bool, | |
| 50 | parse: parser_fn}; | |
| 5829 | 51 | |
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 52 | fun make_parser comment markup int_only parse = | 
| 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 53 |   Parser {comment = comment, markup = markup, int_only = int_only, parse = parse};
 | 
| 5829 | 54 | |
| 55 | ||
| 56 | (* parse command *) | |
| 57 | ||
| 6860 | 58 | local | 
| 6199 | 59 | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 60 | fun terminate false = Scan.succeed () | 
| 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 61 | | 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 | 62 | |
| 26620 | 63 | fun body cmd (name, _) = | 
| 7026 | 64 | (case cmd name of | 
| 24868 | 65 |     SOME (Parser {int_only, parse, ...}) =>
 | 
| 26620 | 66 | P.!!! (Scan.prompt (name ^ "# ") (P.tags |-- parse >> pair int_only)) | 
| 15531 | 67 |   | NONE => sys_error ("no parser for outer syntax command " ^ quote name));
 | 
| 6860 | 68 | |
| 69 | in | |
| 5829 | 70 | |
| 26620 | 71 | fun parse_command do_terminate cmd = | 
| 15531 | 72 | P.semicolon >> K NONE || | 
| 73 | P.sync >> K NONE || | |
| 26620 | 74 | (P.position P.command :-- body cmd) --| terminate do_terminate | 
| 6860 | 75 | >> (fn ((name, pos), (int_only, f)) => | 
| 15531 | 76 | SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |> | 
| 6860 | 77 | Toplevel.interactive int_only |> f)); | 
| 5829 | 78 | |
| 6199 | 79 | end; | 
| 80 | ||
| 5829 | 81 | |
| 82 | ||
| 9132 | 83 | (** global outer syntax **) | 
| 5829 | 84 | |
| 7026 | 85 | local | 
| 86 | ||
| 24868 | 87 | val global_parsers = ref (Symtab.empty: parser Symtab.table); | 
| 22120 | 88 | val global_markups = ref ([]: (string * ThyOutput.markup) list); | 
| 5952 | 89 | |
| 23939 | 90 | fun change_parsers f = CRITICAL (fn () => | 
| 91 | (change global_parsers f; | |
| 92 | global_markups := | |
| 24868 | 93 |     Symtab.fold (fn (name, Parser {markup = SOME m, ...}) => cons (name, m) | _ => I)
 | 
| 94 | (! global_parsers) [])); | |
| 6722 | 95 | |
| 7026 | 96 | in | 
| 97 | ||
| 9132 | 98 | (* access current syntax *) | 
| 7026 | 99 | |
| 24868 | 100 | fun get_parsers () = CRITICAL (fn () => ! global_parsers); | 
| 101 | fun get_markups () = CRITICAL (fn () => ! global_markups); | |
| 7026 | 102 | |
| 24868 | 103 | fun get_parser () = Symtab.lookup (get_parsers ()); | 
| 7789 | 104 | |
| 24868 | 105 | fun is_markup kind name = AList.lookup (op =) (get_markups ()) name = SOME kind; | 
| 5829 | 106 | |
| 107 | ||
| 108 | (* augment syntax *) | |
| 109 | ||
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 110 | fun add_parser name kind parser = CRITICAL (fn () => | 
| 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 111 | (OuterKeyword.command name kind; | 
| 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 112 | if not (Symtab.defined (get_parsers ()) name) then () | 
| 24868 | 113 |   else warning ("Redefining outer syntax command " ^ quote name);
 | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 114 | change_parsers (Symtab.update (name, parser)))); | 
| 5829 | 115 | |
| 24868 | 116 | fun command name comment kind parse = | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 117 | add_parser name kind (make_parser comment NONE false parse); | 
| 5829 | 118 | |
| 24868 | 119 | fun markup_command markup name comment kind parse = | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 120 | add_parser name kind (make_parser comment (SOME markup) false parse); | 
| 24868 | 121 | |
| 122 | fun improper_command name comment kind parse = | |
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 123 | add_parser name kind (make_parser comment NONE true parse); | 
| 7026 | 124 | |
| 125 | end; | |
| 5829 | 126 | |
| 127 | ||
| 26990 | 128 | (* local_theory commands *) | 
| 129 | ||
| 130 | fun local_theory_command do_print trans name comment kind parse = | |
| 131 | command name comment kind (P.opt_target -- parse | |
| 132 | >> (fn (loc, f) => (if do_print then Toplevel.print else I) o trans loc f)); | |
| 133 | ||
| 134 | val local_theory = local_theory_command false Toplevel.local_theory; | |
| 135 | val local_theory_to_proof' = local_theory_command true Toplevel.local_theory_to_proof'; | |
| 136 | val local_theory_to_proof = local_theory_command true Toplevel.local_theory_to_proof; | |
| 137 | ||
| 138 | ||
| 24872 | 139 | (* inspect syntax *) | 
| 7026 | 140 | |
| 141 | fun dest_parsers () = | |
| 16727 | 142 | get_parsers () |> Symtab.dest |> sort_wrt #1 | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 143 |   |> map (fn (name, Parser {comment, int_only, ...}) => (name, comment, int_only));
 | 
| 5829 | 144 | |
| 9223 | 145 | fun print_outer_syntax () = | 
| 7026 | 146 | let | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 147 | fun pretty_cmd (name, comment, _) = | 
| 7026 | 148 | Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 149 | val (int_cmds, cmds) = List.partition #3 (dest_parsers ()); | 
| 7026 | 150 | in | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 151 |     [Pretty.strs ("syntax keywords:" :: map quote (OuterKeyword.dest_keywords ())),
 | 
| 18326 | 152 | Pretty.big_list "commands:" (map pretty_cmd cmds), | 
| 153 | Pretty.big_list "interactive-only commands:" (map pretty_cmd int_cmds)] | |
| 9223 | 154 | |> Pretty.chunks |> Pretty.writeln | 
| 7026 | 155 | end; | 
| 5829 | 156 | |
| 157 | ||
| 158 | ||
| 9132 | 159 | (** toplevel parsing **) | 
| 5829 | 160 | |
| 9132 | 161 | (* basic sources *) | 
| 6860 | 162 | |
| 26620 | 163 | fun toplevel_source term do_recover cmd src = | 
| 9132 | 164 | let | 
| 165 | val no_terminator = | |
| 166 | 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 | 167 | fun recover int = | 
| 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 wenzelm parents: 
23679diff
changeset | 168 | (int, fn _ => Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]); | 
| 9132 | 169 | in | 
| 170 | src | |
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
10749diff
changeset | 171 | |> T.source_proper | 
| 9132 | 172 | |> Source.source T.stopper | 
| 15531 | 173 | (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 | 174 | (Option.map recover do_recover) | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19060diff
changeset | 175 | |> Source.map_filter I | 
| 24868 | 176 | |> Source.source T.stopper | 
| 26620 | 177 | (Scan.bulk (fn xs => P.!!! (parse_command term (cmd ())) xs)) | 
| 23682 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 wenzelm parents: 
23679diff
changeset | 178 | (Option.map recover do_recover) | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19060diff
changeset | 179 | |> Source.map_filter I | 
| 9132 | 180 | end; | 
| 5829 | 181 | |
| 7746 | 182 | |
| 25580 | 183 | (* off-line scanning/parsing *) | 
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 184 | |
| 15144 
85929e1b307d
Remove isar_readstring.  Split read into scanner and parser.
 aspinall parents: 
14981diff
changeset | 185 | fun scan str = | 
| 16195 | 186 | Source.of_string str | 
| 187 | |> Symbol.source false | |
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 188 | |> T.source (SOME false) OuterKeyword.get_lexicons Position.none | 
| 16195 | 189 | |> Source.exhaust; | 
| 190 | ||
| 25580 | 191 | fun parse pos str = | 
| 192 | Source.of_string str | |
| 193 | |> Symbol.source false | |
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 194 | |> T.source (SOME false) OuterKeyword.get_lexicons pos | 
| 26620 | 195 | |> toplevel_source false NONE get_parser | 
| 25580 | 196 | |> Source.exhaust; | 
| 197 | ||
| 14091 | 198 | |
| 26431 | 199 | (* process file *) | 
| 200 | ||
| 201 | fun process_file path thy = | |
| 202 | let | |
| 203 | val result = ref thy; | |
| 26881 | 204 | val trs = parse (Path.position path) (File.read path); | 
| 27575 | 205 | val init = Toplevel.init_theory "" (K thy) (fn thy' => result := thy'); | 
| 26431 | 206 | val _ = Toplevel.excursion (init Toplevel.empty :: trs @ [Toplevel.exit Toplevel.empty]); | 
| 207 | in ! result end; | |
| 208 | ||
| 209 | ||
| 24868 | 210 | (* interactive source of toplevel transformers *) | 
| 211 | ||
| 26600 | 212 | type isar = | 
| 213 | (Toplevel.transition, (Toplevel.transition option, | |
| 214 | (OuterLex.token, (OuterLex.token option, (OuterLex.token, (OuterLex.token, | |
| 215 | Position.T * (Symbol.symbol, (string, unit) Source.source) Source.source) | |
| 216 | Source.source) Source.source) Source.source) Source.source) Source.source) Source.source; | |
| 217 | ||
| 218 | fun isar term : isar = | |
| 24868 | 219 | Source.tty | 
| 220 | |> Symbol.source true | |
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 221 | |> T.source (SOME true) OuterKeyword.get_lexicons Position.none | 
| 26620 | 222 | |> toplevel_source term (SOME true) get_parser; | 
| 24868 | 223 | |
| 224 | ||
| 26611 | 225 | (* load_thy *) | 
| 7746 | 226 | |
| 26323 
73efc70edeef
theory loader: discontinued *attached* ML scripts;
 wenzelm parents: 
26291diff
changeset | 227 | fun load_thy dir name pos text time = | 
| 7683 | 228 | let | 
| 24065 | 229 | val text_src = Source.of_list (Library.untabify text); | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 230 | val (lexs, parser) = CRITICAL (fn () => (OuterKeyword.get_lexicons (), get_parser ())); | 
| 23866 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 231 | |
| 17932 | 232 | val _ = Present.init_theory name; | 
| 24065 | 233 | val _ = Present.verbatim_source name (fn () => Source.exhaust (Symbol.source false text_src)); | 
| 234 | val toks = text_src | |
| 17932 | 235 | |> Symbol.source false | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 236 | |> T.source NONE (K lexs) pos | 
| 17932 | 237 | |> Source.exhausted; | 
| 238 | val trs = toks | |
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 239 | |> toplevel_source false NONE (K parser) | 
| 17932 | 240 | |> Source.exhaust; | 
| 23866 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 241 | |
| 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 242 |     val _ = if time then writeln ("\n**** Starting theory " ^ quote name ^ " ****") else ();
 | 
| 25685 | 243 | val _ = cond_timeit time "" (fn () => | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
26990diff
changeset | 244 | ThyOutput.process_thy (#1 lexs) OuterKeyword.command_tags is_markup trs toks | 
| 23866 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 245 | |> Buffer.content | 
| 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 246 | |> Present.theory_output name); | 
| 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 247 |     val _ = if time then writeln ("**** Finished theory " ^ quote name ^ " ****\n") else ();
 | 
| 24065 | 248 | in () end; | 
| 23866 
5295671034f8
moved deps_thy to ThyLoad (independent of outer syntax);
 wenzelm parents: 
23796diff
changeset | 249 | |
| 5829 | 250 | end; | 
| 251 |