| author | hoelzl | 
| Mon, 20 Oct 2014 18:33:14 +0200 | |
| changeset 58729 | e8ecc79aee43 | 
| parent 57918 | f5d73caba4e5 | 
| child 58850 | 1bb0ad7827b4 | 
| 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 | 
| 55448 
e42a3fc18458
explicit indication that redefining outer syntax commands is not supposed to happen -- NB: interactive mode requires global change of syntax;
 wenzelm parents: 
54734diff
changeset | 13 | val batch_mode: bool Unsynchronized.ref | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43711diff
changeset | 14 | val is_markup: outer_syntax -> Thy_Output.markup -> string -> bool | 
| 43711 | 15 | val get_syntax: unit -> (Scan.lexicon * Scan.lexicon) * outer_syntax | 
| 46970 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 wenzelm parents: 
46961diff
changeset | 16 | val check_syntax: unit -> unit | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48638diff
changeset | 17 | type command_spec = (string * Keyword.T) * Position.T | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 18 | val command: command_spec -> string -> | 
| 29311 | 19 | (Toplevel.transition -> Toplevel.transition) parser -> unit | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 20 | val markup_command: Thy_Output.markup -> command_spec -> string -> | 
| 29311 | 21 | (Toplevel.transition -> Toplevel.transition) parser -> unit | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 22 | val improper_command: command_spec -> string -> | 
| 43711 | 23 | (Toplevel.transition -> Toplevel.transition) parser -> unit | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 24 | val local_theory': command_spec -> string -> | 
| 29380 | 25 | (bool -> local_theory -> local_theory) parser -> unit | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 26 | val local_theory: command_spec -> string -> | 
| 29311 | 27 | (local_theory -> local_theory) parser -> unit | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 28 | val local_theory_to_proof': command_spec -> string -> | 
| 29311 | 29 | (bool -> local_theory -> Proof.state) parser -> unit | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 30 | val local_theory_to_proof: command_spec -> string -> | 
| 29311 | 31 | (local_theory -> Proof.state) parser -> unit | 
| 50213 | 32 | val help_outer_syntax: string list -> unit | 
| 5883 | 33 | val print_outer_syntax: unit -> unit | 
| 57918 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 34 | val scan: Scan.lexicon * Scan.lexicon -> Position.T -> string -> Token.T list | 
| 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 35 | val parse: (Scan.lexicon * Scan.lexicon) * outer_syntax -> | 
| 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 36 | Position.T -> string -> Toplevel.transition list | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 37 | val parse_spans: Token.T list -> Command_Span.span list | 
| 26600 | 38 | type isar | 
| 38253 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
38236diff
changeset | 39 | val isar: TextIO.instream -> bool -> isar | 
| 52510 | 40 | val side_comments: Token.T list -> Token.T list | 
| 57105 | 41 | val command_reports: outer_syntax -> Token.T -> Position.report_text list | 
| 52510 | 42 | val read_spans: outer_syntax -> Token.T list -> Toplevel.transition list | 
| 5829 | 43 | end; | 
| 44 | ||
| 36953 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 wenzelm parents: 
36950diff
changeset | 45 | structure Outer_Syntax: OUTER_SYNTAX = | 
| 5829 | 46 | struct | 
| 47 | ||
| 48 | (** outer syntax **) | |
| 49 | ||
| 29311 | 50 | (* command parsers *) | 
| 5829 | 51 | |
| 29311 | 52 | datatype command = Command of | 
| 24868 | 53 |  {comment: string,
 | 
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36959diff
changeset | 54 | markup: Thy_Output.markup option, | 
| 24868 | 55 | int_only: bool, | 
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 56 | parse: (Toplevel.transition -> Toplevel.transition) parser, | 
| 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 57 | pos: Position.T, | 
| 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 58 | id: serial}; | 
| 5829 | 59 | |
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 60 | fun new_command comment markup int_only parse pos = | 
| 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 61 |   Command {comment = comment, markup = markup, int_only = int_only, parse = parse,
 | 
| 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 62 | pos = pos, id = serial ()}; | 
| 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 63 | |
| 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 64 | fun command_markup def (name, Command {pos, id, ...}) =
 | 
| 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 65 | Markup.properties (Position.entity_properties_of def id pos) | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49564diff
changeset | 66 | (Markup.entity Markup.commandN name); | 
| 5829 | 67 | |
| 50213 | 68 | fun pretty_command (cmd as (name, Command {comment, ...})) =
 | 
| 69 | Pretty.block | |
| 70 | (Pretty.marks_str | |
| 50450 
358b6020f8b6
generalized notion of active area, where sendback is just one application;
 wenzelm parents: 
50215diff
changeset | 71 |       ([Active.make_markup Markup.sendbackN {implicit = true, properties = [Markup.padding_line]},
 | 
| 50215 | 72 | command_markup false cmd], name) :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); | 
| 50213 | 73 | |
| 5829 | 74 | |
| 75 | (* parse command *) | |
| 76 | ||
| 6860 | 77 | local | 
| 6199 | 78 | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 79 | fun terminate false = Scan.succeed () | 
| 44357 | 80 | | terminate true = | 
| 81 | Parse.group (fn () => "end of input") | |
| 82 | (Scan.option Parse.sync -- Parse.semicolon >> K ()); | |
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 83 | |
| 26620 | 84 | fun body cmd (name, _) = | 
| 7026 | 85 | (case cmd name of | 
| 29311 | 86 |     SOME (Command {int_only, parse, ...}) =>
 | 
| 36950 | 87 | Parse.!!! (Scan.prompt (name ^ "# ") (Parse.tags |-- parse >> pair int_only)) | 
| 48191 
c1def7433a72
internalize error into command transaction -- relevant for commands that are declared via 'keywords', but not defined yet;
 wenzelm parents: 
47416diff
changeset | 88 | | NONE => | 
| 
c1def7433a72
internalize error into command transaction -- relevant for commands that are declared via 'keywords', but not defined yet;
 wenzelm parents: 
47416diff
changeset | 89 | Scan.succeed (false, Toplevel.imperative (fn () => | 
| 
c1def7433a72
internalize error into command transaction -- relevant for commands that are declared via 'keywords', but not defined yet;
 wenzelm parents: 
47416diff
changeset | 90 |         error ("Bad parser for outer syntax command " ^ quote name))));
 | 
| 6860 | 91 | |
| 92 | in | |
| 5829 | 93 | |
| 26620 | 94 | fun parse_command do_terminate cmd = | 
| 36950 | 95 | Parse.semicolon >> K NONE || | 
| 96 | Parse.sync >> K NONE || | |
| 97 | (Parse.position Parse.command :-- body cmd) --| terminate do_terminate | |
| 6860 | 98 | >> (fn ((name, pos), (int_only, f)) => | 
| 15531 | 99 | SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |> | 
| 6860 | 100 | Toplevel.interactive int_only |> f)); | 
| 5829 | 101 | |
| 6199 | 102 | end; | 
| 103 | ||
| 5829 | 104 | |
| 43711 | 105 | (* type outer_syntax *) | 
| 106 | ||
| 107 | datatype outer_syntax = Outer_Syntax of | |
| 108 |  {commands: command Symtab.table,
 | |
| 109 | markups: (string * Thy_Output.markup) list}; | |
| 110 | ||
| 111 | fun make_outer_syntax commands markups = | |
| 112 |   Outer_Syntax {commands = commands, markups = markups};
 | |
| 113 | ||
| 114 | val empty_outer_syntax = make_outer_syntax Symtab.empty []; | |
| 115 | ||
| 116 | ||
| 117 | fun map_commands f (Outer_Syntax {commands, ...}) =
 | |
| 118 | let | |
| 119 | val commands' = f commands; | |
| 120 | val markups' = | |
| 121 |       Symtab.fold (fn (name, Command {markup = SOME m, ...}) => cons (name, m) | _ => I)
 | |
| 122 | commands' []; | |
| 123 | in make_outer_syntax commands' markups' end; | |
| 124 | ||
| 125 | fun dest_commands (Outer_Syntax {commands, ...}) =
 | |
| 50213 | 126 | commands |> Symtab.dest |> sort_wrt #1; | 
| 43711 | 127 | |
| 128 | fun lookup_commands (Outer_Syntax {commands, ...}) = Symtab.lookup commands;
 | |
| 129 | ||
| 130 | fun is_markup (Outer_Syntax {markups, ...}) kind name =
 | |
| 131 | AList.lookup (op =) markups name = SOME kind; | |
| 132 | ||
| 133 | ||
| 5829 | 134 | |
| 9132 | 135 | (** global outer syntax **) | 
| 5829 | 136 | |
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48638diff
changeset | 137 | type command_spec = (string * Keyword.T) * Position.T; | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 138 | |
| 55448 
e42a3fc18458
explicit indication that redefining outer syntax commands is not supposed to happen -- NB: interactive mode requires global change of syntax;
 wenzelm parents: 
54734diff
changeset | 139 | val batch_mode = Unsynchronized.ref false; | 
| 
e42a3fc18458
explicit indication that redefining outer syntax commands is not supposed to happen -- NB: interactive mode requires global change of syntax;
 wenzelm parents: 
54734diff
changeset | 140 | |
| 7026 | 141 | local | 
| 142 | ||
| 43711 | 143 | (*synchronized wrt. Keywords*) | 
| 144 | val global_outer_syntax = Unsynchronized.ref empty_outer_syntax; | |
| 5952 | 145 | |
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 146 | fun add_command (name, kind) cmd = CRITICAL (fn () => | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46876diff
changeset | 147 | let | 
| 56294 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
56203diff
changeset | 148 | val context = ML_Context.the_generic_context (); | 
| 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
56203diff
changeset | 149 | val thy = Context.theory_of context; | 
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 150 |     val Command {pos, ...} = cmd;
 | 
| 57623 | 151 | val command_name = quote (Markup.markup Markup.keyword1 name); | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46876diff
changeset | 152 | val _ = | 
| 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46876diff
changeset | 153 | (case try (Thy_Header.the_keyword thy) name of | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 154 | SOME spec => | 
| 48864 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 wenzelm parents: 
48771diff
changeset | 155 | if Option.map #1 spec = SOME (Keyword.kind_files_of kind) then () | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48638diff
changeset | 156 |           else error ("Inconsistent outer syntax keyword declaration " ^
 | 
| 57623 | 157 | command_name ^ Position.here pos) | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46876diff
changeset | 158 | | NONE => | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 159 | if Context.theory_name thy = Context.PureN | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 160 | then Keyword.define (name, SOME kind) | 
| 57623 | 161 |           else error ("Undeclared outer syntax command " ^ command_name ^ Position.here pos));
 | 
| 56294 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
56203diff
changeset | 162 | val _ = Context_Position.report_generic context pos (command_markup true (name, cmd)); | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46876diff
changeset | 163 | in | 
| 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46876diff
changeset | 164 | Unsynchronized.change global_outer_syntax (map_commands (fn commands => | 
| 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46876diff
changeset | 165 | (if not (Symtab.defined commands name) then () | 
| 57026 | 166 | else if ! batch_mode then | 
| 57623 | 167 |         error ("Attempt to redefine outer syntax command " ^ command_name)
 | 
| 55448 
e42a3fc18458
explicit indication that redefining outer syntax commands is not supposed to happen -- NB: interactive mode requires global change of syntax;
 wenzelm parents: 
54734diff
changeset | 168 | else | 
| 57623 | 169 |         warning ("Redefining outer syntax command " ^ command_name ^
 | 
| 57026 | 170 | Position.here (Position.thread_data ())); | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46876diff
changeset | 171 | Symtab.update (name, cmd) commands))) | 
| 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46876diff
changeset | 172 | end); | 
| 6722 | 173 | |
| 7026 | 174 | in | 
| 175 | ||
| 43711 | 176 | fun get_syntax () = CRITICAL (fn () => (Keyword.get_lexicons (), ! global_outer_syntax)); | 
| 7789 | 177 | |
| 46970 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 wenzelm parents: 
46961diff
changeset | 178 | fun check_syntax () = | 
| 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 wenzelm parents: 
46961diff
changeset | 179 | let | 
| 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 wenzelm parents: 
46961diff
changeset | 180 | val ((_, major), syntax) = CRITICAL (fn () => (Keyword.dest (), ! global_outer_syntax)); | 
| 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 wenzelm parents: 
46961diff
changeset | 181 | in | 
| 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 wenzelm parents: 
46961diff
changeset | 182 | (case subtract (op =) (map #1 (dest_commands syntax)) major of | 
| 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 wenzelm parents: 
46961diff
changeset | 183 | [] => () | 
| 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 wenzelm parents: 
46961diff
changeset | 184 |     | missing => error ("Missing outer syntax command(s) " ^ commas_quote missing))
 | 
| 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 wenzelm parents: 
46961diff
changeset | 185 | end; | 
| 
9667e0dcb5e2
check declared vs. defined commands at end of session;
 wenzelm parents: 
46961diff
changeset | 186 | |
| 43711 | 187 | fun lookup_commands_dynamic () = lookup_commands (! global_outer_syntax); | 
| 5829 | 188 | |
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 189 | fun command (spec, pos) comment parse = | 
| 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 190 | add_command spec (new_command comment NONE false parse pos); | 
| 24868 | 191 | |
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 192 | fun markup_command markup (spec, pos) comment parse = | 
| 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 193 | add_command spec (new_command comment (SOME markup) false parse pos); | 
| 7026 | 194 | |
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 195 | fun improper_command (spec, pos) comment parse = | 
| 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 196 | add_command spec (new_command comment NONE true parse pos); | 
| 29311 | 197 | |
| 43711 | 198 | end; | 
| 199 | ||
| 5829 | 200 | |
| 26990 | 201 | (* local_theory commands *) | 
| 202 | ||
| 56895 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 wenzelm parents: 
56334diff
changeset | 203 | fun local_theory_command trans command_spec comment parse = | 
| 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 wenzelm parents: 
56334diff
changeset | 204 | command command_spec comment (Parse.opt_target -- parse >> (fn (loc, f) => trans loc f)); | 
| 26990 | 205 | |
| 56895 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 wenzelm parents: 
56334diff
changeset | 206 | val local_theory' = local_theory_command Toplevel.local_theory'; | 
| 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 wenzelm parents: 
56334diff
changeset | 207 | val local_theory = local_theory_command Toplevel.local_theory; | 
| 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 wenzelm parents: 
56334diff
changeset | 208 | val local_theory_to_proof' = local_theory_command Toplevel.local_theory_to_proof'; | 
| 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 wenzelm parents: 
56334diff
changeset | 209 | val local_theory_to_proof = local_theory_command Toplevel.local_theory_to_proof; | 
| 26990 | 210 | |
| 211 | ||
| 24872 | 212 | (* inspect syntax *) | 
| 7026 | 213 | |
| 50213 | 214 | fun help_outer_syntax pats = | 
| 215 | dest_commands (#2 (get_syntax ())) | |
| 216 | |> filter (fn (name, _) => forall (fn pat => match_string pat name) pats) | |
| 217 | |> map pretty_command | |
| 56334 
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
 wenzelm parents: 
56294diff
changeset | 218 | |> Pretty.writeln_chunks; | 
| 50213 | 219 | |
| 9223 | 220 | fun print_outer_syntax () = | 
| 7026 | 221 | let | 
| 46957 
0c15caf47040
clarified Keyword.is_keyword: union of minor and major;
 wenzelm parents: 
46950diff
changeset | 222 | val ((keywords, _), outer_syntax) = | 
| 
0c15caf47040
clarified Keyword.is_keyword: union of minor and major;
 wenzelm parents: 
46950diff
changeset | 223 | CRITICAL (fn () => (Keyword.dest (), #2 (get_syntax ()))); | 
| 50213 | 224 | val (int_cmds, cmds) = | 
| 225 |       List.partition (fn (_, Command {int_only, ...}) => int_only) (dest_commands outer_syntax);
 | |
| 7026 | 226 | in | 
| 43711 | 227 |     [Pretty.strs ("syntax keywords:" :: map quote keywords),
 | 
| 50213 | 228 | Pretty.big_list "commands:" (map pretty_command cmds), | 
| 229 | Pretty.big_list "interactive-only commands:" (map pretty_command int_cmds)] | |
| 56334 
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
 wenzelm parents: 
56294diff
changeset | 230 | |> Pretty.writeln_chunks | 
| 7026 | 231 | end; | 
| 5829 | 232 | |
| 233 | ||
| 234 | ||
| 9132 | 235 | (** toplevel parsing **) | 
| 5829 | 236 | |
| 9132 | 237 | (* basic sources *) | 
| 6860 | 238 | |
| 26620 | 239 | fun toplevel_source term do_recover cmd src = | 
| 9132 | 240 | let | 
| 241 | 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 | 242 | 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 | 243 | fun recover int = | 
| 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 wenzelm parents: 
23679diff
changeset | 244 | (int, fn _ => Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]); | 
| 9132 | 245 | in | 
| 246 | src | |
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36955diff
changeset | 247 | |> 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 | 248 | |> Source.source Token.stopper | 
| 51627 
589daaf48dba
tuned signature -- agree with markup terminology;
 wenzelm parents: 
51271diff
changeset | 249 | (Scan.bulk (Parse.$$$ "--" -- Parse.!!! Parse.document_source >> K NONE || Parse.not_eof >> SOME)) | 
| 23682 
cf4773532006
nested source: explicit interactive flag for recover avoids duplicate errors;
 wenzelm parents: 
23679diff
changeset | 250 | (Option.map recover do_recover) | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19060diff
changeset | 251 | |> 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 | 252 | |> Source.source Token.stopper | 
| 36950 | 253 | (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 | 254 | (Option.map recover do_recover) | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19060diff
changeset | 255 | |> Source.map_filter I | 
| 9132 | 256 | end; | 
| 5829 | 257 | |
| 7746 | 258 | |
| 25580 | 259 | (* off-line scanning/parsing *) | 
| 14925 
0f86a8a694f8
added read (provides transition names and sources);
 wenzelm parents: 
14687diff
changeset | 260 | |
| 57918 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 261 | fun scan lexs pos = | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 262 | Source.of_string #> | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 263 | Symbol.source #> | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 264 |   Token.source {do_recover = SOME false} (K lexs) pos #>
 | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 265 | Source.exhaust; | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 266 | |
| 57918 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 267 | fun parse (lexs, syntax) pos str = | 
| 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 268 | Source.of_string str | 
| 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 269 | |> Symbol.source | 
| 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 270 |   |> Token.source {do_recover = SOME false} (K lexs) pos
 | 
| 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 271 | |> toplevel_source false NONE (K (lookup_commands syntax)) | 
| 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 272 | |> Source.exhaust; | 
| 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57905diff
changeset | 273 | |
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 274 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 275 | (* parse spans *) | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 276 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 277 | local | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 278 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 279 | fun ship span = | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 280 | let | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 281 | val kind = | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 282 | if not (null span) andalso Token.is_command (hd span) andalso not (exists Token.is_error span) | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 283 | then Command_Span.Command_Span (Token.content_of (hd span), Token.pos_of (hd span)) | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 284 | else if forall Token.is_improper span then Command_Span.Ignored_Span | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 285 | else Command_Span.Malformed_Span; | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 286 | in cons (Command_Span.Span (kind, span)) end; | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 287 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 288 | fun flush (result, content, improper) = | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 289 | result | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 290 | |> not (null content) ? ship (rev content) | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 291 | |> not (null improper) ? ship (rev improper); | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 292 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 293 | fun parse tok (result, content, improper) = | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 294 | if Token.is_command tok then (flush (result, content, improper), [tok], []) | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 295 | else if Token.is_improper tok then (result, content, tok :: improper) | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 296 | else (result, tok :: (improper @ content), []); | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 297 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 298 | in | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 299 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 300 | fun parse_spans toks = | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 301 | fold parse toks ([], [], []) |> flush |> rev; | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 302 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 303 | end; | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 304 | |
| 14091 | 305 | |
| 24868 | 306 | (* interactive source of toplevel transformers *) | 
| 307 | ||
| 26600 | 308 | type isar = | 
| 309 | (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 | 310 | (Token.T, (Token.T option, (Token.T, (Token.T, | 
| 54734 
b91afc3aa3e6
clarified Proof General legacy: special treatment of \<^newline> only in TTY mode;
 wenzelm parents: 
52510diff
changeset | 311 | (Symbol_Pos.T, | 
| 
b91afc3aa3e6
clarified Proof General legacy: special treatment of \<^newline> only in TTY mode;
 wenzelm parents: 
52510diff
changeset | 312 | Position.T * (Symbol.symbol, (Symbol.symbol, (string, unit) Source.source) Source.source) | 
| 27770 | 313 | Source.source) Source.source) Source.source) Source.source) | 
| 314 | Source.source) Source.source) Source.source) Source.source; | |
| 26600 | 315 | |
| 38253 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
38236diff
changeset | 316 | fun isar in_stream term : isar = | 
| 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 wenzelm parents: 
38236diff
changeset | 317 | 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 | 318 | |> Symbol.source | 
| 56203 
76c72f4d0667
clarified bootstrap process: switch to ML with context and antiquotations earlier;
 wenzelm parents: 
55708diff
changeset | 319 | |> Source.map_filter (fn "\<^newline>" => SOME "\n" | s => SOME s) (*Proof General legacy*) | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36955diff
changeset | 320 |   |> Token.source {do_recover = SOME true} Keyword.get_lexicons Position.none
 | 
| 43711 | 321 | |> toplevel_source term (SOME true) lookup_commands_dynamic; | 
| 24868 | 322 | |
| 323 | ||
| 52510 | 324 | (* side-comments *) | 
| 48749 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 wenzelm parents: 
48647diff
changeset | 325 | |
| 52510 | 326 | fun cmts (t1 :: t2 :: toks) = | 
| 327 | if Token.keyword_with (fn s => s = "--") t1 then t2 :: cmts toks | |
| 328 | else cmts (t2 :: toks) | |
| 329 | | cmts _ = []; | |
| 330 | ||
| 331 | val side_comments = filter Token.is_proper #> cmts; | |
| 332 | ||
| 333 | ||
| 334 | (* read commands *) | |
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 335 | |
| 52510 | 336 | fun command_reports outer_syntax tok = | 
| 337 | if Token.is_command tok then | |
| 338 | let val name = Token.content_of tok in | |
| 339 | (case lookup_commands outer_syntax name of | |
| 340 | NONE => [] | |
| 55708 | 341 | | SOME cmd => [((Token.pos_of tok, command_markup false (name, cmd)), "")]) | 
| 52510 | 342 | end | 
| 343 | else []; | |
| 48749 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 wenzelm parents: 
48647diff
changeset | 344 | |
| 52510 | 345 | fun read_spans outer_syntax toks = | 
| 346 | Source.of_list toks | |
| 347 | |> toplevel_source false NONE (K (lookup_commands outer_syntax)) | |
| 348 | |> Source.exhaust; | |
| 28432 
944cb67f809a
load_thy: Toplevel.excursion based on units of command/proof pairs;
 wenzelm parents: 
28424diff
changeset | 349 | |
| 5829 | 350 | end; | 
| 36953 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 wenzelm parents: 
36950diff
changeset | 351 |