| author | huffman | 
| Sat, 13 Mar 2010 18:16:48 -0800 | |
| changeset 35779 | 7de1e14d9277 | 
| parent 35021 | c839a4c670c6 | 
| child 36950 | 75b8f26f2f07 | 
| 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 BASIC_ML_CONTEXT = | |
| 8 | sig | |
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 9 | val bind_thm: string * thm -> unit | 
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 10 | val bind_thms: string * thm list -> unit | 
| 24574 | 11 | val thm: xstring -> thm | 
| 12 | val thms: xstring -> thm list | |
| 25204 | 13 | end | 
| 24574 | 14 | |
| 15 | signature ML_CONTEXT = | |
| 16 | sig | |
| 17 | include BASIC_ML_CONTEXT | |
| 18 | val the_generic_context: unit -> Context.generic | |
| 26432 
095e448b95a0
renamed ML_Context.the_context to ML_Context.the_global_context;
 wenzelm parents: 
26416diff
changeset | 19 | val the_global_context: unit -> theory | 
| 24574 | 20 | val the_local_context: unit -> Proof.context | 
| 26455 | 21 | val exec: (unit -> unit) -> Context.generic -> Context.generic | 
| 32738 | 22 | val stored_thms: thm list Unsynchronized.ref | 
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 23 | val ml_store_thm: string * thm -> unit | 
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 24 | val ml_store_thms: string * thm list -> unit | 
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
33222diff
changeset | 25 | type antiq = Proof.context -> (Proof.context -> string * string) * Proof.context | 
| 30513 | 26 | val add_antiq: string -> (Position.T -> antiq context_parser) -> unit | 
| 32738 | 27 | val trace: bool Unsynchronized.ref | 
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 28 | val eval_antiquotes: ML_Lex.token Antiquote.antiquote list * Position.T -> | 
| 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 29 | Context.generic option -> (ML_Lex.token list * ML_Lex.token list) * Context.generic option | 
| 30574 | 30 | val eval: bool -> Position.T -> Symbol_Pos.text -> unit | 
| 26455 | 31 | val eval_file: Path.T -> unit | 
| 30574 | 32 | val eval_in: Proof.context option -> bool -> Position.T -> Symbol_Pos.text -> unit | 
| 32738 | 33 | val evaluate: Proof.context -> bool -> | 
| 34 | string * (unit -> 'a) option Unsynchronized.ref -> string -> 'a | |
| 26455 | 35 | val expression: Position.T -> string -> string -> string -> Context.generic -> Context.generic | 
| 25204 | 36 | end | 
| 24574 | 37 | |
| 38 | structure ML_Context: ML_CONTEXT = | |
| 39 | struct | |
| 40 | ||
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 41 | (** implicit ML context **) | 
| 24574 | 42 | |
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 43 | 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 | 44 | val the_global_context = Context.theory_of o the_generic_context; | 
| 24574 | 45 | val the_local_context = Context.proof_of o the_generic_context; | 
| 46 | ||
| 26455 | 47 | fun exec (e: unit -> unit) context = | 
| 48 | (case Context.setmp_thread_data (SOME context) (fn () => (e (); Context.thread_data ())) () of | |
| 49 | SOME context' => context' | |
| 50 | | NONE => error "Missing context after execution"); | |
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 51 | |
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 52 | |
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 53 | (* theorem bindings *) | 
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 54 | |
| 33222 | 55 | val stored_thms: thm list Unsynchronized.ref = Unsynchronized.ref []; (* FIXME via context!? *) | 
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 56 | |
| 26492 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 57 | fun ml_store sel (name, ths) = | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 58 | let | 
| 29579 | 59 | val ths' = Context.>>> (Context.map_theory_result | 
| 60 | (PureThy.store_thms (Binding.name name, ths))); | |
| 26492 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 61 | val _ = | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 62 | if name = "" then () | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 63 | else if not (ML_Syntax.is_identifier name) then | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 64 |         error ("Cannot bind theorem(s) " ^ quote name ^ " as ML value")
 | 
| 32966 
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
 wenzelm parents: 
32738diff
changeset | 65 | else setmp_CRITICAL stored_thms ths' (fn () => | 
| 31334 | 66 | ML_Compiler.eval true Position.none | 
| 67 |           (ML_Lex.tokenize ("val " ^ name ^ " = " ^ sel ^ "(! ML_Context.stored_thms);"))) ();
 | |
| 26492 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 68 | in () end; | 
| 24574 | 69 | |
| 26492 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 70 | val ml_store_thms = ml_store ""; | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 71 | fun ml_store_thm (name, th) = ml_store "hd" (name, [th]); | 
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 72 | |
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
35019diff
changeset | 73 | fun bind_thm (name, thm) = ml_store_thm (name, Drule.export_without_context thm); | 
| 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
35019diff
changeset | 74 | fun bind_thms (name, thms) = ml_store_thms (name, map Drule.export_without_context thms); | 
| 24574 | 75 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 76 | fun thm name = ProofContext.get_thm (the_local_context ()) name; | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 77 | fun thms name = ProofContext.get_thms (the_local_context ()) name; | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 78 | |
| 24574 | 79 | |
| 80 | ||
| 81 | (** ML antiquotations **) | |
| 82 | ||
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 83 | (* antiquotation commands *) | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 84 | |
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
33222diff
changeset | 85 | type antiq = Proof.context -> (Proof.context -> string * string) * Proof.context; | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 86 | |
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 87 | local | 
| 24574 | 88 | |
| 32738 | 89 | val global_parsers = | 
| 90 | Unsynchronized.ref (Symtab.empty: (Position.T -> antiq context_parser) Symtab.table); | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 91 | |
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 92 | in | 
| 24574 | 93 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 94 | fun add_antiq name scan = CRITICAL (fn () => | 
| 32738 | 95 | Unsynchronized.change global_parsers (fn tab => | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 96 | (if not (Symtab.defined tab name) then () | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 97 |     else warning ("Redefined ML antiquotation: " ^ quote name);
 | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 98 | Symtab.update (name, scan) tab))); | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 99 | |
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 100 | fun antiquotation src ctxt = | 
| 24574 | 101 | let val ((name, _), pos) = Args.dest_src src in | 
| 102 | (case Symtab.lookup (! global_parsers) name of | |
| 103 |       NONE => error ("Unknown ML antiquotation command: " ^ quote name ^ Position.str_of pos)
 | |
| 28412 
0608c04858c7
back to plain Position.report for regular references;
 wenzelm parents: 
28407diff
changeset | 104 | | SOME scan => (Position.report (Markup.ML_antiq name) pos; | 
| 27895 | 105 | Args.context_syntax "ML antiquotation" (scan pos) src ctxt)) | 
| 24574 | 106 | end; | 
| 107 | ||
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 108 | end; | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 109 | |
| 24574 | 110 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 111 | (* parsing and evaluation *) | 
| 24574 | 112 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 113 | local | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 114 | |
| 24574 | 115 | structure P = OuterParse; | 
| 30588 | 116 | structure T = OuterLex; | 
| 24574 | 117 | |
| 118 | val antiq = | |
| 27809 
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27781diff
changeset | 119 | P.!!! (P.position P.xname -- Args.parse --| Scan.ahead P.eof) | 
| 24574 | 120 | >> (fn ((x, pos), y) => Args.src ((x, y), pos)); | 
| 121 | ||
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
33222diff
changeset | 122 | val begin_env = ML_Lex.tokenize "structure Isabelle =\nstruct\n"; | 
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 123 | val end_env = ML_Lex.tokenize "end;"; | 
| 31334 | 124 | val reset_env = ML_Lex.tokenize "structure Isabelle = struct end"; | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 125 | |
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 126 | in | 
| 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 127 | |
| 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 128 | fun eval_antiquotes (ants, pos) opt_context = | 
| 24574 | 129 | let | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 130 | val opt_ctxt = Option.map (Context.Proof o Context.proof_of) opt_context; | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 131 | val ((ml_env, ml_body), opt_ctxt') = | 
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 132 | if forall Antiquote.is_text ants | 
| 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 133 | then (([], map (fn Antiquote.Text tok => tok) ants), opt_ctxt) | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 134 | else | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 135 | let | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 136 | val ctxt = | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 137 | (case opt_ctxt of | 
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 138 |               NONE => error ("No context -- cannot expand ML antiquotations" ^ Position.str_of pos)
 | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 139 | | SOME ctxt => Context.proof_of ctxt); | 
| 27378 | 140 | |
| 27359 | 141 | val lex = #1 (OuterKeyword.get_lexicons ()); | 
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 142 | fun no_decl _ = ([], []); | 
| 27378 | 143 | |
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 144 | fun expand (Antiquote.Text tok) state = (K ([], [tok]), state) | 
| 30683 
e8ac1f9d9469
datatype antiquote: maintain original Position.range, which is eventually attached to the resulting ML tokens;
 wenzelm parents: 
30672diff
changeset | 145 | | expand (Antiquote.Antiq (ss, range)) (scope, background) = | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 146 | let | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 147 | val context = Stack.top scope; | 
| 30683 
e8ac1f9d9469
datatype antiquote: maintain original Position.range, which is eventually attached to the resulting ML tokens;
 wenzelm parents: 
30672diff
changeset | 148 | val (f, context') = antiquotation (T.read_antiq lex antiq (ss, #1 range)) context; | 
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
33222diff
changeset | 149 | val (decl, background') = f background; | 
| 30683 
e8ac1f9d9469
datatype antiquote: maintain original Position.range, which is eventually attached to the resulting ML tokens;
 wenzelm parents: 
30672diff
changeset | 150 | val decl' = decl #> pairself (ML_Lex.tokenize #> map (ML_Lex.set_range range)); | 
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 151 | in (decl', (Stack.map_top (K context') scope, background')) end | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 152 | | expand (Antiquote.Open _) (scope, background) = | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 153 | (no_decl, (Stack.push scope, background)) | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 154 | | expand (Antiquote.Close _) (scope, background) = | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 155 | (no_decl, (Stack.pop scope, background)); | 
| 24574 | 156 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 157 | val (decls, (_, ctxt')) = fold_map expand ants (Stack.init ctxt, ctxt); | 
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 158 | val ml = decls |> map (fn decl => decl ctxt') |> split_list |> pairself flat; | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 159 | in (ml, SOME (Context.Proof ctxt')) end; | 
| 30637 
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
 wenzelm parents: 
30614diff
changeset | 160 | in ((begin_env @ ml_env @ end_env, ml_body), opt_ctxt') end; | 
| 24574 | 161 | |
| 32738 | 162 | val trace = Unsynchronized.ref false; | 
| 24574 | 163 | |
| 30672 
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
 wenzelm parents: 
30643diff
changeset | 164 | fun eval verbose pos txt = | 
| 24574 | 165 | let | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 166 | (*prepare source text*) | 
| 27878 | 167 | val _ = Position.report Markup.ML_source pos; | 
| 30643 | 168 | val ants = ML_Lex.read_antiq (Symbol_Pos.explode (txt, pos), pos); | 
| 31334 | 169 | val ((env, body), env_ctxt) = eval_antiquotes (ants, pos) (Context.thread_data ()); | 
| 170 | val _ = | |
| 171 | if ! trace then tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body]) | |
| 172 | else (); | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 173 | |
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 174 | (*prepare static ML environment*) | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 175 | val _ = Context.setmp_thread_data env_ctxt | 
| 31334 | 176 | (fn () => (ML_Compiler.eval false Position.none env; Context.thread_data ())) () | 
| 31325 
700951b53d21
moved local ML environment to separate module ML_Env;
 wenzelm parents: 
30683diff
changeset | 177 | |> (fn NONE => () | SOME context' => Context.>> (ML_Env.inherit context')); | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 178 | |
| 31334 | 179 | val _ = ML_Compiler.eval verbose pos body; | 
| 180 | val _ = ML_Compiler.eval false Position.none reset_env; | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 181 | in () end; | 
| 24574 | 182 | |
| 183 | end; | |
| 184 | ||
| 185 | ||
| 30672 
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
 wenzelm parents: 
30643diff
changeset | 186 | (* derived versions *) | 
| 24574 | 187 | |
| 26881 | 188 | fun eval_file path = eval true (Path.position path) (File.read path); | 
| 24574 | 189 | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 190 | fun eval_in ctxt verbose pos txt = | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 191 | Context.setmp_thread_data (Option.map Context.Proof ctxt) (fn () => eval verbose pos txt) (); | 
| 25204 | 192 | |
| 33222 | 193 | (* FIXME not thread-safe -- reference can be accessed directly *) | 
| 30672 
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
 wenzelm parents: 
30643diff
changeset | 194 | fun evaluate ctxt verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () => | 
| 25204 | 195 | let | 
| 196 | val _ = r := NONE; | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 197 | val _ = Context.setmp_thread_data (SOME (Context.Proof ctxt)) (fn () => | 
| 30672 
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
 wenzelm parents: 
30643diff
changeset | 198 |       eval verbose Position.none ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))")) ();
 | 
| 25204 | 199 |   in (case ! r of NONE => error ("Bad evaluation for " ^ ref_name) | SOME e => e) end) ();
 | 
| 24574 | 200 | |
| 26455 | 201 | fun expression pos bind body txt = | 
| 202 | exec (fn () => eval false pos | |
| 203 |     ("Context.set_thread_data (SOME (let " ^ bind ^ " = " ^ txt ^ " in " ^ body ^
 | |
| 204 | " end (ML_Context.the_generic_context ())));")); | |
| 205 | ||
| 24574 | 206 | end; | 
| 207 | ||
| 208 | structure Basic_ML_Context: BASIC_ML_CONTEXT = ML_Context; | |
| 209 | open Basic_ML_Context; |