| author | Fabian Huch <huch@in.tum.de> | 
| Wed, 05 Jun 2024 20:09:04 +0200 | |
| changeset 80257 | 96cb31f0bbdf | 
| parent 78035 | bd5f6cee8001 | 
| child 82587 | 7415414bd9d8 | 
| permissions | -rw-r--r-- | 
| 5829 | 1 | (* Title: Pure/Isar/outer_syntax.ML | 
| 2 | Author: Markus Wenzel, TU Muenchen | |
| 3 | ||
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 4 | Isabelle/Isar outer syntax. | 
| 5829 | 5 | *) | 
| 6 | ||
| 7 | signature OUTER_SYNTAX = | |
| 8 | sig | |
| 58930 | 9 | val help: theory -> string list -> unit | 
| 10 | val print_commands: theory -> unit | |
| 59935 | 11 | type command_keyword = string * Position.T | 
| 12 | val command: command_keyword -> string -> | |
| 29311 | 13 | (Toplevel.transition -> Toplevel.transition) parser -> unit | 
| 63273 | 14 | val maybe_begin_local_theory: command_keyword -> string -> | 
| 15 | (local_theory -> local_theory) parser -> (theory -> local_theory) parser -> unit | |
| 59935 | 16 | val local_theory': command_keyword -> string -> | 
| 29380 | 17 | (bool -> local_theory -> local_theory) parser -> unit | 
| 59935 | 18 | val local_theory: command_keyword -> string -> | 
| 29311 | 19 | (local_theory -> local_theory) parser -> unit | 
| 59935 | 20 | val local_theory_to_proof': command_keyword -> string -> | 
| 29311 | 21 | (bool -> local_theory -> Proof.state) parser -> unit | 
| 59935 | 22 | val local_theory_to_proof: command_keyword -> string -> | 
| 29311 | 23 | (local_theory -> Proof.state) parser -> unit | 
| 60095 | 24 | val bootstrap: bool Config.T | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 25 | val parse_spans: Token.T list -> Command_Span.span list | 
| 68184 
6c693b2700b3
support for dynamic document output while editing;
 wenzelm parents: 
67499diff
changeset | 26 | val make_span: Token.T list -> Command_Span.span | 
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 27 | val parse_span: theory -> (unit -> theory) -> Token.T list -> Toplevel.transition | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 28 | val parse_text: theory -> (unit -> theory) -> Position.T -> string -> Toplevel.transition list | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 29 | val command_reports: theory -> Token.T -> Position.report_text list | 
| 63274 | 30 | val check_command: Proof.context -> command_keyword -> string | 
| 5829 | 31 | end; | 
| 32 | ||
| 36953 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 wenzelm parents: 
36950diff
changeset | 33 | structure Outer_Syntax: OUTER_SYNTAX = | 
| 5829 | 34 | struct | 
| 35 | ||
| 36 | (** outer syntax **) | |
| 37 | ||
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 38 | (* errors *) | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 39 | |
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 40 | fun err_command msg name ps = | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 41 | error (msg ^ quote (Markup.markup Markup.keyword1 name) ^ Position.here_list ps); | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 42 | |
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 43 | fun err_dup_command name ps = | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 44 | err_command "Duplicate outer syntax command " name ps; | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 45 | |
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 46 | |
| 29311 | 47 | (* command parsers *) | 
| 5829 | 48 | |
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 49 | datatype command_parser = | 
| 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 50 | Parser of (Toplevel.transition -> Toplevel.transition) parser | | 
| 60691 | 51 | Restricted_Parser of | 
| 59939 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 wenzelm parents: 
59935diff
changeset | 52 | (bool * Position.T) option -> (Toplevel.transition -> Toplevel.transition) parser; | 
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 53 | |
| 29311 | 54 | datatype command = Command of | 
| 24868 | 55 |  {comment: string,
 | 
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 56 | command_parser: command_parser, | 
| 48647 
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 | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 60 | fun eq_command (Command {id = id1, ...}, Command {id = id2, ...}) = id1 = id2;
 | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 61 | |
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 62 | fun new_command comment command_parser pos = | 
| 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 63 |   Command {comment = comment, command_parser = command_parser, pos = pos, id = serial ()};
 | 
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 64 | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 65 | fun command_pos (Command {pos, ...}) = pos;
 | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 66 | |
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 67 | fun command_markup def (name, Command {pos, id, ...}) =
 | 
| 74183 | 68 | Position.make_entity_markup def id Markup.commandN (name, pos); | 
| 5829 | 69 | |
| 50213 | 70 | fun pretty_command (cmd as (name, Command {comment, ...})) =
 | 
| 71 | Pretty.block | |
| 72 | (Pretty.marks_str | |
| 50450 
358b6020f8b6
generalized notion of active area, where sendback is just one application;
 wenzelm parents: 
50215diff
changeset | 73 |       ([Active.make_markup Markup.sendbackN {implicit = true, properties = [Markup.padding_line]},
 | 
| 74262 | 74 |         command_markup {def = false} cmd], name) :: Pretty.str ":" ::
 | 
| 75 | Pretty.brk 2 :: Pretty.text comment); | |
| 50213 | 76 | |
| 5829 | 77 | |
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 78 | (* theory data *) | 
| 43711 | 79 | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 80 | structure Data = Theory_Data | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 81 | ( | 
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 82 | type T = command Symtab.table; | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 83 | val empty = Symtab.empty; | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 84 | fun merge data : T = | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 85 | data |> Symtab.join (fn name => fn (cmd1, cmd2) => | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 86 | if eq_command (cmd1, cmd2) then raise Symtab.SAME | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 87 | else err_dup_command name [command_pos cmd1, command_pos cmd2]); | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 88 | ); | 
| 43711 | 89 | |
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 90 | val get_commands = Data.get; | 
| 60924 
610794dff23c
tuned signature, in accordance to sortBy in Scala;
 wenzelm parents: 
60693diff
changeset | 91 | val dest_commands = get_commands #> Symtab.dest #> sort_by #1; | 
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 92 | val lookup_commands = Symtab.lookup o get_commands; | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 93 | |
| 58930 | 94 | fun help thy pats = | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 95 | dest_commands thy | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 96 | |> filter (fn (name, _) => forall (fn pat => match_string pat name) pats) | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 97 | |> map pretty_command | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 98 | |> Pretty.writeln_chunks; | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 99 | |
| 58930 | 100 | fun print_commands thy = | 
| 43711 | 101 | let | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 102 | val keywords = Thy_Header.get_keywords thy; | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 103 | val minor = Scan.dest_lexicon (Keyword.minor_keywords keywords); | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 104 | val commands = dest_commands thy; | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 105 | in | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 106 |     [Pretty.strs ("keywords:" :: map quote minor),
 | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 107 | Pretty.big_list "commands:" (map pretty_command commands)] | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 108 | |> Pretty.writeln_chunks | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 109 | end; | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 110 | |
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 111 | |
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 112 | (* maintain commands *) | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 113 | |
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 114 | fun add_command name cmd thy = | 
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
76815diff
changeset | 115 | if member (op =) Thy_Header.bootstrap_thys (Context.theory_base_name thy) then thy | 
| 63022 | 116 | else | 
| 117 | let | |
| 118 | val _ = | |
| 119 | Keyword.is_command (Thy_Header.get_keywords thy) name orelse | |
| 120 | err_command "Undeclared outer syntax command " name [command_pos cmd]; | |
| 121 | val _ = | |
| 122 | (case lookup_commands thy name of | |
| 123 | NONE => () | |
| 124 | | SOME cmd' => err_dup_command name [command_pos cmd, command_pos cmd']); | |
| 125 | val _ = | |
| 126 | Context_Position.report_generic (Context.the_generic_context ()) | |
| 74262 | 127 |           (command_pos cmd) (command_markup {def = true} (name, cmd));
 | 
| 63022 | 128 | in Data.map (Symtab.update (name, cmd)) thy end; | 
| 43711 | 129 | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 130 | val _ = Theory.setup (Theory.at_end (fn thy => | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 131 | let | 
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 132 | val command_keywords = | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 133 | Scan.dest_lexicon (Keyword.major_keywords (Thy_Header.get_keywords thy)); | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 134 | val _ = | 
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58934diff
changeset | 135 | (case subtract (op =) (map #1 (dest_commands thy)) command_keywords of | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 136 | [] => () | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 137 |       | missing => error ("Missing outer syntax command(s) " ^ commas_quote missing))
 | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 138 | in NONE end)); | 
| 43711 | 139 | |
| 5829 | 140 | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 141 | (* implicit theory setup *) | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46958diff
changeset | 142 | |
| 59935 | 143 | type command_keyword = string * Position.T; | 
| 5952 | 144 | |
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 145 | fun raw_command (name, pos) comment command_parser = | 
| 59932 | 146 | let val setup = add_command name (new_command comment command_parser pos) | 
| 147 | in Context.>> (Context.mapping setup (Local_Theory.background_theory setup)) end; | |
| 26990 | 148 | |
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 149 | fun command (name, pos) comment parse = | 
| 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 150 | raw_command (name, pos) comment (Parser parse); | 
| 59923 
b21c82422d65
support private scope for individual local theory commands;
 wenzelm parents: 
59083diff
changeset | 151 | |
| 63273 | 152 | fun maybe_begin_local_theory command_keyword comment parse_local parse_global = | 
| 153 | raw_command command_keyword comment | |
| 154 | (Restricted_Parser (fn restricted => | |
| 155 | Parse.opt_target -- parse_local | |
| 156 | >> (fn (target, f) => Toplevel.local_theory restricted target f) || | |
| 157 | (if is_some restricted then Scan.fail | |
| 72434 | 158 | else parse_global >> Toplevel.begin_main_target true))); | 
| 63273 | 159 | |
| 59935 | 160 | fun local_theory_command trans command_keyword comment parse = | 
| 161 | raw_command command_keyword comment | |
| 60691 | 162 | (Restricted_Parser (fn restricted => | 
| 163 | Parse.opt_target -- parse >> (fn (target, f) => trans restricted target f))); | |
| 26990 | 164 | |
| 74964 
77a96ed74340
allow general command transactions with presentation;
 wenzelm parents: 
74561diff
changeset | 165 | val local_theory' = | 
| 76815 | 166 | local_theory_command (fn a => fn b => fn c => Toplevel.local_theory' a b c NONE); | 
| 56895 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 wenzelm parents: 
56334diff
changeset | 167 | 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 | 168 | 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 | 169 | val local_theory_to_proof = local_theory_command Toplevel.local_theory_to_proof; | 
| 26990 | 170 | |
| 171 | ||
| 5829 | 172 | |
| 9132 | 173 | (** toplevel parsing **) | 
| 5829 | 174 | |
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 175 | (* parse spans *) | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 176 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 177 | local | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 178 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 179 | fun ship span = | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 180 | let | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 181 | val kind = | 
| 68729 
3a02b424d5fb
clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
 wenzelm parents: 
68184diff
changeset | 182 | if forall Token.is_ignored span then Command_Span.Ignored_Span | 
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 183 | else if exists Token.is_error span then Command_Span.Malformed_Span | 
| 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 184 | else | 
| 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 185 | (case find_first Token.is_command span of | 
| 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 186 | NONE => Command_Span.Malformed_Span | 
| 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 187 | | SOME cmd => Command_Span.Command_Span (Token.content_of cmd, Token.pos_of cmd)); | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 188 | in cons (Command_Span.Span (kind, span)) end; | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 189 | |
| 68729 
3a02b424d5fb
clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
 wenzelm parents: 
68184diff
changeset | 190 | fun flush (result, content, ignored) = | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 191 | result | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 192 | |> not (null content) ? ship (rev content) | 
| 68729 
3a02b424d5fb
clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
 wenzelm parents: 
68184diff
changeset | 193 | |> not (null ignored) ? ship (rev ignored); | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 194 | |
| 68729 
3a02b424d5fb
clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
 wenzelm parents: 
68184diff
changeset | 195 | fun parse tok (result, content, ignored) = | 
| 
3a02b424d5fb
clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
 wenzelm parents: 
68184diff
changeset | 196 | if Token.is_ignored tok then (result, content, tok :: ignored) | 
| 59939 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 wenzelm parents: 
59935diff
changeset | 197 | else if Token.is_command_modifier tok orelse | 
| 59924 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 wenzelm parents: 
59923diff
changeset | 198 | Token.is_command tok andalso | 
| 59939 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 wenzelm parents: 
59935diff
changeset | 199 | (not (exists Token.is_command_modifier content) orelse exists Token.is_command content) | 
| 68729 
3a02b424d5fb
clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
 wenzelm parents: 
68184diff
changeset | 200 | then (flush (result, content, ignored), [tok], []) | 
| 
3a02b424d5fb
clarified ignored span / core range: include formal comments, e.g. relevant for error messages from antiquotations;
 wenzelm parents: 
68184diff
changeset | 201 | else (result, tok :: (ignored @ content), []); | 
| 57905 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 202 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 203 | in | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 204 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 205 | fun parse_spans toks = | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 206 | fold parse toks ([], [], []) |> flush |> rev; | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 207 | |
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 208 | end; | 
| 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 wenzelm parents: 
57623diff
changeset | 209 | |
| 68184 
6c693b2700b3
support for dynamic document output while editing;
 wenzelm parents: 
67499diff
changeset | 210 | fun make_span toks = | 
| 
6c693b2700b3
support for dynamic document output while editing;
 wenzelm parents: 
67499diff
changeset | 211 | (case parse_spans toks of | 
| 
6c693b2700b3
support for dynamic document output while editing;
 wenzelm parents: 
67499diff
changeset | 212 | [span] => span | 
| 
6c693b2700b3
support for dynamic document output while editing;
 wenzelm parents: 
67499diff
changeset | 213 | | _ => Command_Span.Span (Command_Span.Malformed_Span, toks)); | 
| 
6c693b2700b3
support for dynamic document output while editing;
 wenzelm parents: 
67499diff
changeset | 214 | |
| 14091 | 215 | |
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 216 | (* parse commands *) | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 217 | |
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 218 | val bootstrap = Config.declare_bool ("Outer_Syntax.bootstrap", \<^here>) (K true);
 | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 219 | |
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 220 | local | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 221 | |
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 222 | val before_command = | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 223 | Scan.option (Parse.position (Parse.private >> K true || Parse.qualified >> K false)); | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 224 | |
| 73106 
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
 wenzelm parents: 
73098diff
changeset | 225 | fun parse_command thy markers = | 
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 226 | Scan.ahead (before_command |-- Parse.position Parse.command) :|-- (fn (name, pos) => | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 227 | let | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 228 | val keywords = Thy_Header.get_keywords thy; | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 229 | val tr0 = | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 230 | Toplevel.empty | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 231 | |> Toplevel.name name | 
| 73106 
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
 wenzelm parents: 
73098diff
changeset | 232 | |> Toplevel.position pos | 
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 233 | |> Keyword.is_proof_open keywords name ? Toplevel.skip_proof_open | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 234 | |> Keyword.is_proof_close keywords name ? Toplevel.skip_proof_close; | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 235 | val command_markers = | 
| 70134 | 236 | Parse.command |-- Document_Source.old_tags >> | 
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 237 | (fn tags => Toplevel.markers (map Document_Marker.legacy_tag tags @ markers) tr0); | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 238 | in | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 239 | (case lookup_commands thy name of | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 240 |         SOME (Command {command_parser = Parser parse, ...}) =>
 | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 241 | Parse.!!! (command_markers -- parse) >> (op |>) | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 242 |       | SOME (Command {command_parser = Restricted_Parser parse, ...}) =>
 | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 243 | before_command :|-- (fn restricted => | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 244 | Parse.!!! (command_markers -- parse restricted)) >> (op |>) | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 245 | | NONE => | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 246 | Scan.fail_with (fn _ => fn _ => | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 247 | let | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 248 | val msg = | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 249 | if Config.get_global thy bootstrap | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 250 | then "missing theory context for command " | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 251 | else "undefined command "; | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 252 | in msg ^ quote (Markup.markup Markup.keyword1 name) end)) | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 253 | end); | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 254 | |
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 255 | in | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 256 | |
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 257 | fun parse_span thy init span = | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 258 | let | 
| 73098 
8a20737e4ebf
support more command positions, analogous to Command.core_range in Isabelle/Scala;
 wenzelm parents: 
72434diff
changeset | 259 | val full_range = Token.range_of span; | 
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 260 | val core_range = Token.core_range_of span; | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 261 | |
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 262 | val markers = map Token.input_of (filter Token.is_document_marker span); | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 263 | fun parse () = | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 264 | filter Token.is_proper span | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 265 | |> Source.of_list | 
| 73106 
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
 wenzelm parents: 
73098diff
changeset | 266 | |> Source.source Token.stopper (Scan.bulk (fn xs => Parse.!!! (parse_command thy markers) xs)) | 
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 267 | |> Source.exhaust; | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 268 | in | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 269 | (case parse () of | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 270 | [tr] => Toplevel.modify_init init tr | 
| 73106 
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
 wenzelm parents: 
73098diff
changeset | 271 | | [] => Toplevel.ignored (#1 full_range) | 
| 
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
 wenzelm parents: 
73098diff
changeset | 272 | | _ => Toplevel.malformed (#1 core_range) "Exactly one command expected") | 
| 
3df45de0c079
discontinued body_range (again): does not quite work, because Position.thread_data is plain Toplevel.pos_of only;
 wenzelm parents: 
73098diff
changeset | 273 | handle ERROR msg => Toplevel.malformed (#1 core_range) msg | 
| 69891 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 274 | end; | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 275 | |
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 276 | fun parse_text thy init pos text = | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 277 | Symbol_Pos.explode (text, pos) | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 278 |   |> Token.tokenize (Thy_Header.get_keywords thy) {strict = false}
 | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 279 | |> parse_spans | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 280 | |> map (Command_Span.content #> parse_span thy init); | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 281 | |
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 282 | end; | 
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 283 | |
| 
def3ec9cdb7e
document markers are formal comments, and may thus occur anywhere in the command-span;
 wenzelm parents: 
69887diff
changeset | 284 | |
| 61618 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 285 | (* check commands *) | 
| 48647 
a5144c4c26a2
report commands as formal entities, with def/ref positions;
 wenzelm parents: 
48646diff
changeset | 286 | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 287 | fun command_reports thy tok = | 
| 52510 | 288 | if Token.is_command tok then | 
| 289 | let val name = Token.content_of tok in | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58918diff
changeset | 290 | (case lookup_commands thy name of | 
| 52510 | 291 | NONE => [] | 
| 74262 | 292 |       | SOME cmd => [((Token.pos_of tok, command_markup {def = false} (name, cmd)), "")])
 | 
| 52510 | 293 | end | 
| 294 | else []; | |
| 48749 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 wenzelm parents: 
48647diff
changeset | 295 | |
| 61618 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 296 | fun check_command ctxt (name, pos) = | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 297 | let | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 298 | val thy = Proof_Context.theory_of ctxt; | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 299 | val keywords = Thy_Header.get_keywords thy; | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 300 | in | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 301 | if Keyword.is_command keywords name then | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 302 | let | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 303 | val markup = | 
| 67495 | 304 | Token.explode0 keywords name | 
| 61618 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 305 | |> maps (command_reports thy) | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 306 | |> map (#2 o #1); | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 307 | val _ = Context_Position.reports ctxt (map (pair pos) markup); | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 308 | in name end | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 309 | else | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 310 | let | 
| 69289 | 311 | val completion_report = | 
| 312 | Completion.make_report (name, pos) | |
| 61618 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 313 | (fn completed => | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 314 | Keyword.dest_commands keywords | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 315 | |> filter completed | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 316 | |> sort_strings | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 317 | |> map (fn a => (a, (Markup.commandN, a)))); | 
| 69289 | 318 |       in error ("Bad command " ^ quote name ^ Position.here pos ^ completion_report) end
 | 
| 61618 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 319 | end; | 
| 
27af754f50ca
more thorough check_command, including completion;
 wenzelm parents: 
61579diff
changeset | 320 | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
61618diff
changeset | 321 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
61618diff
changeset | 322 | (* 'ML' command -- required for bootstrapping Isar *) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
61618diff
changeset | 323 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
61618diff
changeset | 324 | val _ = | 
| 64556 | 325 |   command ("ML", \<^here>) "ML text within theory or local theory"
 | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
61618diff
changeset | 326 | (Parse.ML_source >> (fn source => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
61618diff
changeset | 327 | Toplevel.generic_theory | 
| 78035 | 328 | (Local_Theory.touch_ml_env #> | 
| 329 | ML_Context.exec (fn () => | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
61618diff
changeset | 330 | ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source) #> | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
61618diff
changeset | 331 | Local_Theory.propagate_ml_env))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
61618diff
changeset | 332 | |
| 5829 | 333 | end; |