| author | wenzelm | 
| Tue, 28 Aug 2018 11:40:11 +0200 | |
| changeset 68829 | 1a4fa494a4a8 | 
| parent 68823 | 5e7b1ae10eb8 | 
| child 68865 | dd44e31ca2c6 | 
| 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 | 
| 56205 | 12 | type decl = Proof.context -> string * string | 
| 59112 | 13 | val value_decl: string -> string -> Proof.context -> decl * Proof.context | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57834diff
changeset | 14 | val add_antiquotation: binding -> (Token.src -> Proof.context -> decl * Proof.context) -> | 
| 56205 | 15 | theory -> theory | 
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59127diff
changeset | 16 | val print_antiquotations: bool -> Proof.context -> unit | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 17 | 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 | 18 | val eval_file: ML_Compiler.flags -> Path.T -> unit | 
| 59064 | 19 | 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 | 20 | 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 | 21 | ML_Lex.token Antiquote.antiquote list -> unit | 
| 59064 | 22 | val eval_source_in: Proof.context option -> ML_Compiler.flags -> Input.source -> unit | 
| 62913 | 23 | val exec: (unit -> unit) -> Context.generic -> Context.generic | 
| 58991 
92b6f4e68c5a
more careful ML source positions, for improved PIDE markup;
 wenzelm parents: 
58979diff
changeset | 24 | val expression: Position.range -> string -> string -> string -> | 
| 58978 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58928diff
changeset | 25 | ML_Lex.token Antiquote.antiquote list -> Context.generic -> Context.generic | 
| 25204 | 26 | end | 
| 24574 | 27 | |
| 28 | structure ML_Context: ML_CONTEXT = | |
| 29 | struct | |
| 30 | ||
| 31 | (** ML antiquotations **) | |
| 32 | ||
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 33 | (* names for generated environment *) | 
| 59112 | 34 | |
| 35 | structure Names = Proof_Data | |
| 36 | ( | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 37 | type T = string * Name.context; | 
| 62900 
c641bf9402fd
simplified default print_depth: context is usually available, in contrast to 0d295e339f52;
 wenzelm parents: 
62889diff
changeset | 38 | 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 | 39 |   fun init _ = ("Isabelle0", init_names);
 | 
| 59112 | 40 | ); | 
| 41 | ||
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 42 | fun struct_name ctxt = #1 (Names.get ctxt); | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 43 | 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 | 44 | |
| 59112 | 45 | fun variant a ctxt = | 
| 46 | let | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 47 | val names = #2 (Names.get ctxt); | 
| 59112 | 48 | 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 | 49 | val ctxt' = (Names.map o apsnd) (K names') ctxt; | 
| 59112 | 50 | in (b, ctxt') end; | 
| 51 | ||
| 52 | ||
| 53 | (* decl *) | |
| 54 | ||
| 55 | type decl = Proof.context -> string * string; (*final context -> ML env, ML body*) | |
| 56 | ||
| 57 | fun value_decl a s ctxt = | |
| 58 | let | |
| 59 | val (b, ctxt') = variant a ctxt; | |
| 60 | val env = "val " ^ b ^ " = " ^ s ^ ";\n"; | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 61 | val body = struct_name ctxt ^ "." ^ b; | 
| 59112 | 62 | fun decl (_: Proof.context) = (env, body); | 
| 63 | in (decl, ctxt') end; | |
| 64 | ||
| 65 | ||
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 66 | (* theory data *) | 
| 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 | ( | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57834diff
changeset | 70 | type T = (Token.src -> Proof.context -> decl * Proof.context) 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 | val extend = I; | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42360diff
changeset | 73 | 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 | 74 | ); | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 75 | |
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 76 | 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 | 77 | |
| 56070 | 78 | fun check_antiquotation ctxt = | 
| 79 | #1 o Name_Space.check (Context.Proof ctxt) (get_antiquotations ctxt); | |
| 80 | ||
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 81 | fun add_antiquotation name f thy = thy | 
| 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 82 | |> Antiquotations.map (Name_Space.define (Context.Theory thy) true (name, f) #> snd); | 
| 55743 | 83 | |
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59127diff
changeset | 84 | 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 | 85 | Pretty.big_list "ML antiquotations:" | 
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59127diff
changeset | 86 | (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 | 87 | |> Pretty.writeln; | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 88 | |
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 89 | fun apply_antiquotation src ctxt = | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61596diff
changeset | 90 | let val (src', f) = Token.check_src ctxt get_antiquotations src | 
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 91 | in f src' ctxt end; | 
| 43563 | 92 | |
| 24574 | 93 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 94 | (* parsing and evaluation *) | 
| 24574 | 95 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 96 | local | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 97 | |
| 24574 | 98 | val antiq = | 
| 63671 | 99 | Parse.!!! ((Parse.token Parse.liberal_name ::: Parse.args) --| Scan.ahead Parse.eof); | 
| 24574 | 100 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 101 | fun make_env name visible = | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 102 | (ML_Lex.tokenize | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 103 |     ("structure " ^ name ^ " =\nstruct\n\
 | 
| 52663 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 104 | \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 | 105 | " (Context.the_local_context ());\n"), | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 106 | 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 | 107 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 108 | 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 | 109 | |
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 110 | fun eval_antiquotes (ants, pos) opt_context = | 
| 24574 | 111 | let | 
| 52663 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 112 | val visible = | 
| 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 113 | (case opt_context of | 
| 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 114 | SOME (Context.Proof ctxt) => Context_Position.is_visible ctxt | 
| 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 115 | | _ => true); | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 116 | val opt_ctxt = Option.map Context.proof_of opt_context; | 
| 52663 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 117 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 118 | val ((ml_env, ml_body), opt_ctxt') = | 
| 61596 | 119 | if forall (fn Antiquote.Text _ => true | _ => false) ants | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 120 | then (([], map (fn Antiquote.Text tok => tok) ants), opt_ctxt) | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 121 | else | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 122 | let | 
| 59112 | 123 | fun tokenize range = apply2 (ML_Lex.tokenize #> map (ML_Lex.set_range range)); | 
| 53167 | 124 | |
| 61471 | 125 | fun expand_src range src ctxt = | 
| 126 | let val (decl, ctxt') = apply_antiquotation src ctxt | |
| 127 | in (decl #> tokenize range, ctxt') end; | |
| 128 | ||
| 61596 | 129 | fun expand (Antiquote.Text tok) ctxt = (K ([], [tok]), ctxt) | 
| 61473 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 wenzelm parents: 
61471diff
changeset | 130 |             | expand (Antiquote.Control {name, range, body}) ctxt =
 | 
| 61595 | 131 | expand_src range | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61596diff
changeset | 132 | (Token.make_src name (if null body then [] else [Token.read_cartouche body])) ctxt | 
| 61473 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 wenzelm parents: 
61471diff
changeset | 133 |             | expand (Antiquote.Antiq {range, body, ...}) ctxt =
 | 
| 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 wenzelm parents: 
61471diff
changeset | 134 | expand_src range | 
| 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 wenzelm parents: 
61471diff
changeset | 135 | (Token.read_antiq (Thy_Header.get_keywords' ctxt) antiq (body, #1 range)) ctxt; | 
| 53167 | 136 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 137 | val ctxt = | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 138 | (case opt_ctxt of | 
| 48992 | 139 |               NONE => error ("No context -- cannot expand ML antiquotations" ^ Position.here pos)
 | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 140 | | SOME ctxt => struct_begin ctxt); | 
| 27378 | 141 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 142 | val (begin_env, end_env) = make_env (struct_name ctxt) visible; | 
| 53167 | 143 | val (decls, ctxt') = fold_map expand ants ctxt; | 
| 52663 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 144 | val (ml_env, ml_body) = | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58991diff
changeset | 145 | decls |> map (fn decl => decl ctxt') |> split_list |> apply2 flat; | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 146 | in ((begin_env @ ml_env @ end_env, ml_body), SOME ctxt') end; | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 147 | in ((ml_env, ml_body), opt_ctxt') end; | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 148 | |
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 149 | in | 
| 24574 | 150 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 151 | fun eval flags pos ants = | 
| 24574 | 152 | 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 | 153 | 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 | 154 | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 155 | (*prepare source text*) | 
| 62889 | 156 | val ((env, body), env_ctxt) = eval_antiquotes (ants, pos) (Context.get_generic_context ()); | 
| 31334 | 157 | val _ = | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 158 | (case env_ctxt of | 
| 41485 
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
 wenzelm parents: 
41375diff
changeset | 159 | SOME ctxt => | 
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56294diff
changeset | 160 | 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 | 161 | then tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body]) | 
| 41375 | 162 | else () | 
| 163 | | NONE => ()); | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 164 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 165 | (*prepare environment*) | 
| 41485 
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
 wenzelm parents: 
41375diff
changeset | 166 | val _ = | 
| 62889 | 167 | Context.setmp_generic_context | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 168 | (Option.map (Context.Proof o Context_Position.set_visible false) env_ctxt) | 
| 62889 | 169 | (fn () => | 
| 170 | (ML_Compiler.eval non_verbose Position.none env; Context.get_generic_context ())) () | |
| 31325 
700951b53d21
moved local ML environment to separate module ML_Env;
 wenzelm parents: 
30683diff
changeset | 171 | |> (fn NONE => () | SOME context' => Context.>> (ML_Env.inherit context')); | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 172 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 173 | (*eval body*) | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 174 | val _ = ML_Compiler.eval flags pos body; | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 175 | |
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 176 | (*clear environment*) | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 177 | val _ = | 
| 62889 | 178 | (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 | 179 | (SOME ctxt, true) => | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 180 | let | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 181 | val name = struct_name ctxt; | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 182 | 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 | 183 | val _ = Context.>> (ML_Env.forget_structure name); | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 184 | in () end | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 185 | | _ => ()); | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 186 | in () end; | 
| 24574 | 187 | |
| 188 | end; | |
| 189 | ||
| 190 | ||
| 30672 
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
 wenzelm parents: 
30643diff
changeset | 191 | (* derived versions *) | 
| 24574 | 192 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 193 | fun eval_file flags path = | 
| 56203 
76c72f4d0667
clarified bootstrap process: switch to ML with context and antiquotations earlier;
 wenzelm parents: 
56201diff
changeset | 194 | let val pos = Path.position path | 
| 68823 | 195 | 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 | 196 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 197 | fun eval_source flags source = | 
| 68820 
2e4df245754e
clarified environment: allow "read>write" specification;
 wenzelm parents: 
68816diff
changeset | 198 | let | 
| 
2e4df245754e
clarified environment: allow "read>write" specification;
 wenzelm parents: 
68816diff
changeset | 199 | val opt_context = Context.get_generic_context (); | 
| 68821 
877534be1930
explicit setup of operations: avoid hardwired stuff;
 wenzelm parents: 
68820diff
changeset | 200 |     val {read_source, ...} = ML_Env.operations opt_context (#environment flags);
 | 
| 
877534be1930
explicit setup of operations: avoid hardwired stuff;
 wenzelm parents: 
68820diff
changeset | 201 | 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 | 202 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 203 | fun eval_in ctxt flags pos ants = | 
| 62889 | 204 | 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 | 205 | (fn () => eval flags pos ants) (); | 
| 24574 | 206 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 207 | fun eval_source_in ctxt flags source = | 
| 62889 | 208 | 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 | 209 | (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 | 210 | |
| 62913 | 211 | fun exec (e: unit -> unit) context = | 
| 212 | (case Context.setmp_generic_context (SOME context) | |
| 213 | (fn () => (e (); Context.get_generic_context ())) () of | |
| 214 | SOME context' => context' | |
| 215 | | NONE => error "Missing context after execution"); | |
| 216 | ||
| 58991 
92b6f4e68c5a
more careful ML source positions, for improved PIDE markup;
 wenzelm parents: 
58979diff
changeset | 217 | fun expression range name constraint body ants = | 
| 59067 | 218 | exec (fn () => | 
| 219 | eval ML_Compiler.flags (#1 range) | |
| 62889 | 220 | (ML_Lex.read "Context.put_generic_context (SOME (let val " @ ML_Lex.read_set_range range name @ | 
| 59067 | 221 |       ML_Lex.read (": " ^ constraint ^ " =") @ ants @
 | 
| 62876 | 222 |       ML_Lex.read ("in " ^ body ^ " end (Context.the_generic_context ())));")));
 | 
| 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 | 223 | |
| 24574 | 224 | end; | 
| 225 | ||
| 62931 | 226 | val ML = ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags); |