| author | wenzelm | 
| Tue, 10 Feb 2015 16:46:21 +0100 | |
| changeset 59499 | 14095f771781 | 
| parent 59127 | 723b11f8ffbf | 
| child 59917 | 9830c944670f | 
| 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 | |
| 9 | val the_generic_context: unit -> Context.generic | |
| 26432 
095e448b95a0
renamed ML_Context.the_context to ML_Context.the_global_context;
 wenzelm parents: 
26416diff
changeset | 10 | val the_global_context: unit -> theory | 
| 24574 | 11 | val the_local_context: unit -> Proof.context | 
| 39164 
e7e12555e763
ML_Context.thm and ML_Context.thms no longer pervasive;
 wenzelm parents: 
39140diff
changeset | 12 | val thm: xstring -> thm | 
| 
e7e12555e763
ML_Context.thm and ML_Context.thms no longer pervasive;
 wenzelm parents: 
39140diff
changeset | 13 | val thms: xstring -> thm list | 
| 26455 | 14 | val exec: (unit -> unit) -> Context.generic -> Context.generic | 
| 56070 | 15 | 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 | 16 | val struct_name: Proof.context -> string | 
| 59112 | 17 | val variant: string -> Proof.context -> string * Proof.context | 
| 56205 | 18 | type decl = Proof.context -> string * string | 
| 59112 | 19 | 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 | 20 | val add_antiquotation: binding -> (Token.src -> Proof.context -> decl * Proof.context) -> | 
| 56205 | 21 | theory -> theory | 
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 22 | val print_antiquotations: Proof.context -> unit | 
| 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 | 
| 58991 
92b6f4e68c5a
more careful ML source positions, for improved PIDE markup;
 wenzelm parents: 
58979diff
changeset | 29 | 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 | 30 | ML_Lex.token Antiquote.antiquote list -> Context.generic -> Context.generic | 
| 25204 | 31 | end | 
| 24574 | 32 | |
| 33 | structure ML_Context: ML_CONTEXT = | |
| 34 | struct | |
| 35 | ||
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 36 | (** implicit ML context **) | 
| 24574 | 37 | |
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 38 | val the_generic_context = Context.the_thread_data; | 
| 26432 
095e448b95a0
renamed ML_Context.the_context to ML_Context.the_global_context;
 wenzelm parents: 
26416diff
changeset | 39 | val the_global_context = Context.theory_of o the_generic_context; | 
| 24574 | 40 | val the_local_context = Context.proof_of o the_generic_context; | 
| 41 | ||
| 42360 | 42 | fun thm name = Proof_Context.get_thm (the_local_context ()) name; | 
| 43 | fun thms name = Proof_Context.get_thms (the_local_context ()) name; | |
| 39164 
e7e12555e763
ML_Context.thm and ML_Context.thms no longer pervasive;
 wenzelm parents: 
39140diff
changeset | 44 | |
| 26455 | 45 | fun exec (e: unit -> unit) context = | 
| 46 | (case Context.setmp_thread_data (SOME context) (fn () => (e (); Context.thread_data ())) () of | |
| 47 | SOME context' => context' | |
| 48 | | NONE => error "Missing context after execution"); | |
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 49 | |
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 50 | |
| 24574 | 51 | |
| 52 | (** ML antiquotations **) | |
| 53 | ||
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 54 | (* names for generated environment *) | 
| 59112 | 55 | |
| 56 | structure Names = Proof_Data | |
| 57 | ( | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 58 | type T = string * Name.context; | 
| 59112 | 59 | val init_names = ML_Syntax.reserved |> fold Name.declare ["ML_context", "ML_print_depth"]; | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 60 |   fun init _ = ("Isabelle0", init_names);
 | 
| 59112 | 61 | ); | 
| 62 | ||
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 63 | fun struct_name ctxt = #1 (Names.get ctxt); | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 64 | 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 | 65 | |
| 59112 | 66 | fun variant a ctxt = | 
| 67 | let | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 68 | val names = #2 (Names.get ctxt); | 
| 59112 | 69 | 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 | 70 | val ctxt' = (Names.map o apsnd) (K names') ctxt; | 
| 59112 | 71 | in (b, ctxt') end; | 
| 72 | ||
| 73 | ||
| 74 | (* decl *) | |
| 75 | ||
| 76 | type decl = Proof.context -> string * string; (*final context -> ML env, ML body*) | |
| 77 | ||
| 78 | fun value_decl a s ctxt = | |
| 79 | let | |
| 80 | val (b, ctxt') = variant a ctxt; | |
| 81 | val env = "val " ^ b ^ " = " ^ s ^ ";\n"; | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 82 | val body = struct_name ctxt ^ "." ^ b; | 
| 59112 | 83 | fun decl (_: Proof.context) = (env, body); | 
| 84 | in (decl, ctxt') end; | |
| 85 | ||
| 86 | ||
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 87 | (* theory data *) | 
| 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 | structure Antiquotations = Theory_Data | 
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42360diff
changeset | 90 | ( | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57834diff
changeset | 91 | 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 | 92 | 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 | 93 | val extend = I; | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42360diff
changeset | 94 | 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 | 95 | ); | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 96 | |
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 97 | 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 | 98 | |
| 56070 | 99 | fun check_antiquotation ctxt = | 
| 100 | #1 o Name_Space.check (Context.Proof ctxt) (get_antiquotations ctxt); | |
| 101 | ||
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 102 | 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 | 103 | |> Antiquotations.map (Name_Space.define (Context.Theory thy) true (name, f) #> snd); | 
| 55743 | 104 | |
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 105 | fun print_antiquotations ctxt = | 
| 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 106 | Pretty.big_list "ML antiquotations:" | 
| 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 107 | (map (Pretty.mark_str o #1) (Name_Space.markup_table ctxt (get_antiquotations ctxt))) | 
| 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 108 | |> Pretty.writeln; | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 109 | |
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 110 | fun apply_antiquotation src ctxt = | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57834diff
changeset | 111 | let val (src', f) = Token.check_src ctxt (get_antiquotations ctxt) src | 
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 112 | in f src' ctxt end; | 
| 43563 | 113 | |
| 24574 | 114 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 115 | (* parsing and evaluation *) | 
| 24574 | 116 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 117 | local | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 118 | |
| 24574 | 119 | val antiq = | 
| 56201 | 120 | Parse.!!! (Parse.position Parse.xname -- Parse.args --| Scan.ahead Parse.eof) | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57834diff
changeset | 121 | >> uncurry Token.src; | 
| 24574 | 122 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 123 | fun make_env name visible = | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 124 | (ML_Lex.tokenize | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 125 |     ("structure " ^ name ^ " =\nstruct\n\
 | 
| 52663 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 126 | \val ML_context = Context_Position.set_visible " ^ Bool.toString visible ^ | 
| 57832 | 127 | " (ML_Context.the_local_context ());\n\ | 
| 57834 
0d295e339f52
prefer dynamic ML_print_depth if context happens to be available;
 wenzelm parents: 
57832diff
changeset | 128 | \val ML_print_depth =\n\ | 
| 
0d295e339f52
prefer dynamic ML_print_depth if context happens to be available;
 wenzelm parents: 
57832diff
changeset | 129 | \ let val default = ML_Options.get_print_depth ()\n\ | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 130 | \ in fn () => ML_Options.get_print_depth_default default end;\n"), | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 131 | 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 | 132 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 133 | 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 | 134 | |
| 59112 | 135 | fun expanding (Antiquote.Text tok) = ML_Lex.is_cartouche tok | 
| 136 | | expanding (Antiquote.Antiq _) = true; | |
| 137 | ||
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 138 | fun eval_antiquotes (ants, pos) opt_context = | 
| 24574 | 139 | let | 
| 52663 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 140 | val visible = | 
| 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 141 | (case opt_context of | 
| 
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
 wenzelm parents: 
50201diff
changeset | 142 | 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 | 143 | | _ => true); | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 144 | 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 | 145 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 146 | val ((ml_env, ml_body), opt_ctxt') = | 
| 59112 | 147 | if forall (not o expanding) ants | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 148 | then (([], map (fn Antiquote.Text tok => tok) ants), opt_ctxt) | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 149 | else | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 150 | let | 
| 59112 | 151 | fun tokenize range = apply2 (ML_Lex.tokenize #> map (ML_Lex.set_range range)); | 
| 53167 | 152 | |
| 59112 | 153 |           fun expand (Antiquote.Antiq (ss, {range, ...})) ctxt =
 | 
| 53167 | 154 | let | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58903diff
changeset | 155 | val keywords = Thy_Header.get_keywords' ctxt; | 
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 156 | val (decl, ctxt') = | 
| 58903 | 157 | apply_antiquotation (Token.read_antiq keywords antiq (ss, #1 range)) ctxt; | 
| 59112 | 158 | in (decl #> tokenize range, ctxt') end | 
| 159 | | expand (Antiquote.Text tok) ctxt = | |
| 160 | if ML_Lex.is_cartouche tok then | |
| 161 | let | |
| 162 | val range = ML_Lex.range_of tok; | |
| 163 | val text = | |
| 164 | Symbol_Pos.explode (ML_Lex.content_of tok, #1 range) | |
| 165 | |> Symbol_Pos.cartouche_content | |
| 166 | |> Symbol_Pos.implode_range range |> #1; | |
| 167 | val (decl, ctxt') = | |
| 168 | value_decl "input" | |
| 169 |                         ("Input.source true " ^ ML_Syntax.print_string text  ^ " " ^
 | |
| 170 | ML_Syntax.atomic (ML_Syntax.print_range range)) ctxt; | |
| 171 | in (decl #> tokenize range, ctxt') end | |
| 172 | else (K ([], [tok]), ctxt); | |
| 53167 | 173 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 174 | val ctxt = | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 175 | (case opt_ctxt of | 
| 48992 | 176 |               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 | 177 | | SOME ctxt => struct_begin ctxt); | 
| 27378 | 178 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 179 | val (begin_env, end_env) = make_env (struct_name ctxt) visible; | 
| 53167 | 180 | 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 | 181 | val (ml_env, ml_body) = | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58991diff
changeset | 182 | 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 | 183 | 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 | 184 | in ((ml_env, ml_body), opt_ctxt') end; | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 185 | |
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 186 | in | 
| 24574 | 187 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 188 | fun eval flags pos ants = | 
| 24574 | 189 | 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 | 190 | 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 | 191 | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 192 | (*prepare source text*) | 
| 41485 
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
 wenzelm parents: 
41375diff
changeset | 193 | val ((env, body), env_ctxt) = eval_antiquotes (ants, pos) (Context.thread_data ()); | 
| 31334 | 194 | val _ = | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 195 | (case env_ctxt of | 
| 41485 
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
 wenzelm parents: 
41375diff
changeset | 196 | SOME ctxt => | 
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56294diff
changeset | 197 | 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 | 198 | then tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body]) | 
| 41375 | 199 | else () | 
| 200 | | NONE => ()); | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 201 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 202 | (*prepare environment*) | 
| 41485 
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
 wenzelm parents: 
41375diff
changeset | 203 | val _ = | 
| 
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
 wenzelm parents: 
41375diff
changeset | 204 | Context.setmp_thread_data | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 205 | (Option.map (Context.Proof o Context_Position.set_visible false) env_ctxt) | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 206 | (fn () => (ML_Compiler.eval non_verbose Position.none env; Context.thread_data ())) () | 
| 31325 
700951b53d21
moved local ML environment to separate module ML_Env;
 wenzelm parents: 
30683diff
changeset | 207 | |> (fn NONE => () | SOME context' => Context.>> (ML_Env.inherit context')); | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 208 | |
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 209 | (*eval body*) | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 210 | val _ = ML_Compiler.eval flags pos body; | 
| 59127 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 211 | |
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 212 | (*clear environment*) | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 213 | val _ = | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 214 | (case (env_ctxt, is_some (Context.thread_data ())) of | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 215 | (SOME ctxt, true) => | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 216 | let | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 217 | val name = struct_name ctxt; | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 218 | 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 | 219 | val _ = Context.>> (ML_Env.forget_structure name); | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 220 | in () end | 
| 
723b11f8ffbf
more careful handling of auxiliary environment structure -- allow nested ML evaluation;
 wenzelm parents: 
59112diff
changeset | 221 | | _ => ()); | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 222 | in () end; | 
| 24574 | 223 | |
| 224 | end; | |
| 225 | ||
| 226 | ||
| 30672 
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
 wenzelm parents: 
30643diff
changeset | 227 | (* derived versions *) | 
| 24574 | 228 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 229 | fun eval_file flags path = | 
| 56203 
76c72f4d0667
clarified bootstrap process: switch to ML with context and antiquotations earlier;
 wenzelm parents: 
56201diff
changeset | 230 | let val pos = Path.position path | 
| 59067 | 231 | in eval flags pos (ML_Lex.read_pos pos (File.read path)) end; | 
| 56203 
76c72f4d0667
clarified bootstrap process: switch to ML with context and antiquotations earlier;
 wenzelm parents: 
56201diff
changeset | 232 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 233 | fun eval_source flags source = | 
| 59064 | 234 | eval flags (Input.pos_of source) (ML_Lex.read_source (#SML 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 | 235 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 236 | fun eval_in ctxt flags pos ants = | 
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 237 | Context.setmp_thread_data (Option.map Context.Proof ctxt) | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 238 | (fn () => eval flags pos ants) (); | 
| 24574 | 239 | |
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 240 | fun eval_source_in ctxt flags source = | 
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
56032diff
changeset | 241 | Context.setmp_thread_data (Option.map Context.Proof ctxt) | 
| 56275 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 wenzelm parents: 
56229diff
changeset | 242 | (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 | 243 | |
| 58991 
92b6f4e68c5a
more careful ML source positions, for improved PIDE markup;
 wenzelm parents: 
58979diff
changeset | 244 | fun expression range name constraint body ants = | 
| 59067 | 245 | exec (fn () => | 
| 246 | eval ML_Compiler.flags (#1 range) | |
| 247 | (ML_Lex.read "Context.set_thread_data (SOME (let val " @ ML_Lex.read_set_range range name @ | |
| 248 |       ML_Lex.read (": " ^ constraint ^ " =") @ ants @
 | |
| 249 |       ML_Lex.read ("in " ^ body ^ " end (ML_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 | 250 | |
| 24574 | 251 | end; | 
| 252 |