| author | wenzelm | 
| Fri, 08 Dec 2023 14:35:24 +0100 | |
| changeset 79201 | 5d27271701a2 | 
| parent 78804 | d4e9d6b7f48d | 
| permissions | -rw-r--r-- | 
| 24581 | 1 | (* Title: Pure/ML/ml_context.ML | 
| 24574 | 2 | Author: Makarius | 
| 3 | ||
| 4 | ML context and antiquotations. | |
| 5 | *) | |
| 6 | ||
| 7 | signature ML_CONTEXT = | |
| 8 | sig | |
| 56070 | 9 | val check_antiquotation: Proof.context -> xstring * Position.T -> string | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 10 | val struct_name: Proof.context -> string | 
| 59112 | 11 | val variant: string -> Proof.context -> string * Proof.context | 
| 73549 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 12 | type decl = Proof.context -> ML_Lex.token list * ML_Lex.token list | 
| 74562 | 13 | type decls = Proof.context -> (ML_Lex.token list * ML_Lex.token list) list | 
| 78804 | 14 | type 'a antiquotation = Position.range -> 'a -> Proof.context -> decl * Proof.context | 
| 15 | val add_antiquotation: binding -> bool -> Token.src antiquotation -> theory -> theory | |
| 16 | val add_antiquotation_embedded: binding -> Input.source antiquotation -> theory -> theory | |
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59127diff
changeset | 17 | val print_antiquotations: bool -> Proof.context -> unit | 
| 73551 | 18 | val expand_antiquotes: ML_Lex.token Antiquote.antiquote list -> | 
| 19 | Proof.context -> decl * Proof.context | |
| 74562 | 20 | val expand_antiquotes_list: ML_Lex.token Antiquote.antiquote list list -> | 
| 21 | Proof.context -> decls * Proof.context | |
| 78689 | 22 | val read_antiquotes: Input.source -> Proof.context -> decl * Proof.context | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 23 | val eval: ML_Compiler.flags -> Position.T -> ML_Lex.token Antiquote.antiquote list -> unit | 
| 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 24 | val eval_file: ML_Compiler.flags -> Path.T -> unit | 
| 59064 | 25 | val eval_source: ML_Compiler.flags -> Input.source -> unit | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 26 | val eval_in: Proof.context option -> ML_Compiler.flags -> Position.T -> | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
36959diff
changeset | 27 | ML_Lex.token Antiquote.antiquote list -> unit | 
| 59064 | 28 | val eval_source_in: Proof.context option -> ML_Compiler.flags -> Input.source -> unit | 
| 62913 | 29 | val exec: (unit -> unit) -> Context.generic -> Context.generic | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 30 | val expression: Position.T -> ML_Lex.token Antiquote.antiquote list -> | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 31 | Context.generic -> Context.generic | 
| 25204 | 32 | end | 
| 24574 | 33 | |
| 34 | structure ML_Context: ML_CONTEXT = | |
| 35 | struct | |
| 36 | ||
| 37 | (** ML antiquotations **) | |
| 38 | ||
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 39 | (* names for generated environment *) | 
| 59112 | 40 | |
| 41 | structure Names = Proof_Data | |
| 42 | ( | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 43 | type T = string * Name.context; | 
| 62900 
c641bf9402fd
simplified default print_depth: context is usually available, in contrast to 0d295e339f52;
 wenzelm parents: 
62889diff
changeset | 44 | val init_names = ML_Syntax.reserved |> Name.declare "ML_context"; | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 45 |   fun init _ = ("Isabelle0", init_names);
 | 
| 59112 | 46 | ); | 
| 47 | ||
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 48 | fun struct_name ctxt = #1 (Names.get ctxt); | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 49 | val struct_begin = (Names.map o apfst) (fn _ => "Isabelle" ^ serial_string ()); | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 50 | |
| 59112 | 51 | fun variant a ctxt = | 
| 52 | let | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 53 | val names = #2 (Names.get ctxt); | 
| 59112 | 54 | val (b, names') = Name.variant (Name.desymbolize (SOME false) a) names; | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 55 | val ctxt' = (Names.map o apsnd) (K names') ctxt; | 
| 59112 | 56 | in (b, ctxt') end; | 
| 57 | ||
| 58 | ||
| 73549 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 59 | (* theory data *) | 
| 59112 | 60 | |
| 73549 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 61 | type decl = | 
| 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 62 | Proof.context -> ML_Lex.token list * ML_Lex.token list; (*final context -> ML env, ML body*) | 
| 59112 | 63 | |
| 74562 | 64 | type decls = Proof.context -> (ML_Lex.token list * ML_Lex.token list) list; | 
| 65 | ||
| 78804 | 66 | type 'a antiquotation = Position.range -> 'a -> Proof.context -> decl * Proof.context; | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 67 | |
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 68 | structure Antiquotations = Theory_Data | 
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42360diff
changeset | 69 | ( | 
| 78804 | 70 | type T = (bool * Token.src antiquotation) Name_Space.table; | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
48992diff
changeset | 71 | val empty : T = Name_Space.empty_table Markup.ML_antiquotationN; | 
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42360diff
changeset | 72 | fun merge data : T = Name_Space.merge_tables data; | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42360diff
changeset | 73 | ); | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 74 | |
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 75 | val get_antiquotations = Antiquotations.get o Proof_Context.theory_of; | 
| 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 76 | |
| 56070 | 77 | fun check_antiquotation ctxt = | 
| 78 | #1 o Name_Space.check (Context.Proof ctxt) (get_antiquotations ctxt); | |
| 79 | ||
| 73549 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 80 | fun add_antiquotation name embedded antiquotation thy = | 
| 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 81 | thy |> Antiquotations.map | 
| 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 82 | (Name_Space.define (Context.Theory thy) true (name, (embedded, antiquotation)) #> #2); | 
| 55743 | 83 | |
| 78804 | 84 | fun add_antiquotation_embedded name antiquotation = | 
| 85 | add_antiquotation name true | |
| 86 | (fn range => fn src => fn ctxt => | |
| 87 | antiquotation range (Token.read ctxt Parse.embedded_input src) ctxt); | |
| 88 | ||
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59127diff
changeset | 89 | fun print_antiquotations verbose ctxt = | 
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 90 | Pretty.big_list "ML antiquotations:" | 
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59127diff
changeset | 91 | (map (Pretty.mark_str o #1) (Name_Space.markup_table verbose ctxt (get_antiquotations ctxt))) | 
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 92 | |> Pretty.writeln; | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 93 | |
| 73549 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 94 | fun apply_antiquotation range src ctxt = | 
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69216diff
changeset | 95 | let | 
| 73549 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 96 | val (src', (embedded, antiquotation)) = Token.check_src ctxt get_antiquotations src; | 
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69216diff
changeset | 97 | val _ = | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69216diff
changeset | 98 | if Options.default_bool "update_control_cartouches" then | 
| 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69216diff
changeset | 99 | Context_Position.reports_text ctxt | 
| 73549 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 100 | (Antiquote.update_reports embedded (#1 range) (map Token.content_of src')) | 
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69216diff
changeset | 101 | else (); | 
| 73549 
a2c589d5e1e4
clarified signature: more detailed token positions for antiquotations;
 wenzelm parents: 
69592diff
changeset | 102 | in antiquotation range src' ctxt end; | 
| 43563 | 103 | |
| 24574 | 104 | |
| 73551 | 105 | (* parsing *) | 
| 24574 | 106 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 107 | local | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 108 | |
| 24574 | 109 | val antiq = | 
| 63671 | 110 | Parse.!!! ((Parse.token Parse.liberal_name ::: Parse.args) --| Scan.ahead Parse.eof); | 
| 24574 | 111 | |
| 73551 | 112 | fun expand_antiquote ant ctxt = | 
| 113 | (case ant of | |
| 114 | Antiquote.Text tok => (K ([], [tok]), ctxt) | |
| 115 |   | Antiquote.Control {name, range, body} =>
 | |
| 116 | ctxt |> apply_antiquotation range | |
| 117 | (Token.make_src name (if null body then [] else [Token.read_cartouche body])) | |
| 118 |   | Antiquote.Antiq {range, body, ...} =>
 | |
| 119 | ctxt |> apply_antiquotation range | |
| 74564 | 120 | (Parse.read_antiq (Thy_Header.get_keywords' ctxt) antiq (body, #1 range))); | 
| 73551 | 121 | |
| 122 | in | |
| 123 | ||
| 124 | fun expand_antiquotes ants ctxt = | |
| 125 | let | |
| 126 | val (decls, ctxt') = fold_map expand_antiquote ants ctxt; | |
| 127 | fun decl ctxt'' = decls |> map (fn d => d ctxt'') |> split_list |> apply2 flat; | |
| 128 | in (decl, ctxt') end; | |
| 129 | ||
| 74562 | 130 | fun expand_antiquotes_list antss ctxt = | 
| 131 | let | |
| 132 | val (decls, ctxt') = fold_map expand_antiquotes antss ctxt; | |
| 133 | fun decl' ctxt'' = map (fn decl => decl ctxt'') decls; | |
| 134 | in (decl', ctxt') end | |
| 135 | ||
| 73551 | 136 | end; | 
| 137 | ||
| 78689 | 138 | val read_antiquotes = expand_antiquotes o ML_Lex.read_source; | 
| 139 | ||
| 73551 | 140 | |
| 141 | (* evaluation *) | |
| 142 | ||
| 143 | local | |
| 144 | ||
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 145 | fun make_env name visible = | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 146 | (ML_Lex.tokenize | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 147 |     ("structure " ^ name ^ " =\nstruct\n\
 | 
| 52663 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 148 | \val ML_context = Context_Position.set_visible " ^ Bool.toString visible ^ | 
| 62900 
c641bf9402fd
simplified default print_depth: context is usually available, in contrast to 0d295e339f52;
 wenzelm parents: 
62889diff
changeset | 149 | " (Context.the_local_context ());\n"), | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 150 | ML_Lex.tokenize "end;"); | 
| 48776 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
47005diff
changeset | 151 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 152 | fun reset_env name = ML_Lex.tokenize ("structure " ^ name ^ " = struct end");
 | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 153 | |
| 73551 | 154 | fun eval_antiquotes ants opt_context = | 
| 155 | if forall Antiquote.is_text ants orelse is_none opt_context then | |
| 156 | (([], map (Antiquote.the_text "No context -- cannot expand ML antiquotations") ants), | |
| 157 | Option.map Context.proof_of opt_context) | |
| 158 | else | |
| 159 | let | |
| 160 | val context = the opt_context; | |
| 161 | val visible = | |
| 162 | (case context of | |
| 163 | Context.Proof ctxt => Context_Position.is_visible ctxt | |
| 164 | | _ => true); | |
| 165 | val ctxt = struct_begin (Context.proof_of context); | |
| 166 | val (begin_env, end_env) = make_env (struct_name ctxt) visible; | |
| 167 | val (decl, ctxt') = expand_antiquotes ants ctxt; | |
| 168 | val (ml_env, ml_body) = decl ctxt'; | |
| 169 | in ((begin_env @ ml_env @ end_env, ml_body), SOME ctxt') end; | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 170 | |
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 171 | in | 
| 24574 | 172 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 173 | fun eval flags pos ants = | 
| 24574 | 174 | let | 
| 56304 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 wenzelm parents: 
56303diff
changeset | 175 | val non_verbose = ML_Compiler.verbose false flags; | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 176 | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 177 | (*prepare source text*) | 
| 73551 | 178 | val ((env, body), env_ctxt) = eval_antiquotes ants (Context.get_generic_context ()); | 
| 31334 | 179 | val _ = | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 180 | (case env_ctxt of | 
| 41485 
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
 wenzelm parents: 
41375diff
changeset | 181 | SOME ctxt => | 
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56294diff
changeset | 182 | if Config.get ctxt ML_Options.source_trace andalso Context_Position.is_visible ctxt | 
| 56294 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
56279diff
changeset | 183 | then tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body]) | 
| 41375 | 184 | else () | 
| 185 | | NONE => ()); | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 186 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 187 | (*prepare environment*) | 
| 41485 
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
 wenzelm parents: 
41375diff
changeset | 188 | val _ = | 
| 62889 | 189 | Context.setmp_generic_context | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 190 | (Option.map (Context.Proof o Context_Position.set_visible false) env_ctxt) | 
| 62889 | 191 | (fn () => | 
| 192 | (ML_Compiler.eval non_verbose Position.none env; Context.get_generic_context ())) () | |
| 68865 
dd44e31ca2c6
support multiple inheritance of ML environments, with canonical merge order as in Context.begin_theory;
 wenzelm parents: 
68823diff
changeset | 193 | |> (fn NONE => () | SOME context' => Context.>> (ML_Env.inherit [context'])); | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 194 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 195 | (*eval body*) | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 196 | val _ = ML_Compiler.eval flags pos body; | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 197 | |
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 198 | (*clear environment*) | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 199 | val _ = | 
| 62889 | 200 | (case (env_ctxt, is_some (Context.get_generic_context ())) of | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 201 | (SOME ctxt, true) => | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 202 | let | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 203 | val name = struct_name ctxt; | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 204 | val _ = ML_Compiler.eval non_verbose Position.none (reset_env name); | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 205 | val _ = Context.>> (ML_Env.forget_structure name); | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 206 | in () end | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 207 | | _ => ()); | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 208 | in () end; | 
| 24574 | 209 | |
| 210 | end; | |
| 211 | ||
| 212 | ||
| 30672 
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
 wenzelm parents: 
30643diff
changeset | 213 | (* derived versions *) | 
| 24574 | 214 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 215 | fun eval_file flags path = | 
| 76884 | 216 | let val pos = Position.file (File.symbolic_path path) | 
| 68823 | 217 | in eval flags pos (ML_Lex.read_text (File.read path, pos)) end; | 
| 56203 
76c72f4d0667
clarified bootstrap process: switch to ML with context and antiquotations earlier;
 wenzelm parents: 
56201diff
changeset | 218 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 219 | fun eval_source flags source = | 
| 68820 
2e4df245754e
clarified environment: allow "read>write" specification;
 wenzelm parents: 
68816diff
changeset | 220 | let | 
| 
2e4df245754e
clarified environment: allow "read>write" specification;
 wenzelm parents: 
68816diff
changeset | 221 | val opt_context = Context.get_generic_context (); | 
| 68821 
877534be1930
explicit setup of operations: avoid hardwired stuff;
 wenzelm parents: 
68820diff
changeset | 222 |     val {read_source, ...} = ML_Env.operations opt_context (#environment flags);
 | 
| 
877534be1930
explicit setup of operations: avoid hardwired stuff;
 wenzelm parents: 
68820diff
changeset | 223 | in eval flags (Input.pos_of source) (read_source source) end; | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
36959diff
changeset | 224 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 225 | fun eval_in ctxt flags pos ants = | 
| 62889 | 226 | Context.setmp_generic_context (Option.map Context.Proof ctxt) | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 227 | (fn () => eval flags pos ants) (); | 
| 24574 | 228 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 229 | fun eval_source_in ctxt flags source = | 
| 62889 | 230 | Context.setmp_generic_context (Option.map Context.Proof ctxt) | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 231 | (fn () => eval_source flags source) (); | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
36959diff
changeset | 232 | |
| 62913 | 233 | fun exec (e: unit -> unit) context = | 
| 234 | (case Context.setmp_generic_context (SOME context) | |
| 235 | (fn () => (e (); Context.get_generic_context ())) () of | |
| 236 | SOME context' => context' | |
| 237 | | NONE => error "Missing context after execution"); | |
| 238 | ||
| 78035 | 239 | fun expression pos ants = | 
| 240 | Local_Theory.touch_ml_env #> | |
| 241 | exec (fn () => eval ML_Compiler.flags pos ants); | |
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
36959diff
changeset | 242 | |
| 24574 | 243 | end; | 
| 244 | ||
| 62931 | 245 | val ML = ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags); | 
| 78728 
72631efa3821
explicitly reject 'handle' with catch-all patterns;
 wenzelm parents: 
78689diff
changeset | 246 | val ML_command = ML_Context.eval_source ML_Compiler.flags; |