| author | wenzelm | 
| Mon, 16 Feb 2009 20:07:05 +0100 | |
| changeset 29751 | e2756594c414 | 
| parent 29606 | fedb8be05f24 | 
| child 30513 | 1796b8ea88aa | 
| 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 | 
| 28279 | 22 | val inherit_env: Proof.context -> Proof.context -> Proof.context | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 23 | val name_space: ML_NameSpace.nameSpace | 
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 24 | val stored_thms: thm list ref | 
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 25 | val ml_store_thm: string * thm -> unit | 
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 26 | val ml_store_thms: string * thm list -> unit | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 27 | type antiq = | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 28 |     {struct_name: string, background: Proof.context} ->
 | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 29 | (Proof.context -> string * string) * Proof.context | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 30 | val add_antiq: string -> | 
| 27868 | 31 | (Position.T -> Context.generic * Args.T list -> antiq * (Context.generic * Args.T list)) -> unit | 
| 24574 | 32 | val trace: bool ref | 
| 27723 | 33 | val eval_wrapper: (string -> unit) * (string -> 'a) -> bool -> Position.T -> string -> unit | 
| 26455 | 34 | val eval: bool -> Position.T -> string -> unit | 
| 35 | val eval_file: Path.T -> unit | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 36 | val eval_in: Proof.context option -> bool -> Position.T -> string -> unit | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 37 | val evaluate: Proof.context -> (string -> unit) * (string -> 'b) -> bool -> | 
| 25204 | 38 | string * (unit -> 'a) option ref -> string -> 'a | 
| 26455 | 39 | val expression: Position.T -> string -> string -> string -> Context.generic -> Context.generic | 
| 25204 | 40 | end | 
| 24574 | 41 | |
| 42 | structure ML_Context: ML_CONTEXT = | |
| 43 | struct | |
| 44 | ||
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 45 | (** implicit ML context **) | 
| 24574 | 46 | |
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 47 | 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 | 48 | val the_global_context = Context.theory_of o the_generic_context; | 
| 24574 | 49 | val the_local_context = Context.proof_of o the_generic_context; | 
| 50 | ||
| 26455 | 51 | fun exec (e: unit -> unit) context = | 
| 52 | (case Context.setmp_thread_data (SOME context) (fn () => (e (); Context.thread_data ())) () of | |
| 53 | SOME context' => context' | |
| 54 | | NONE => error "Missing context after execution"); | |
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 55 | |
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 56 | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 57 | (* ML name space *) | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 58 | |
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 59 | structure ML_Env = GenericDataFun | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 60 | ( | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 61 | type T = | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 62 | ML_NameSpace.valueVal Symtab.table * | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 63 | ML_NameSpace.typeVal Symtab.table * | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 64 | ML_NameSpace.fixityVal Symtab.table * | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 65 | ML_NameSpace.structureVal Symtab.table * | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 66 | ML_NameSpace.signatureVal Symtab.table * | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 67 | ML_NameSpace.functorVal Symtab.table; | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 68 | val empty = (Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty); | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 69 | val extend = I; | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 70 | fun merge _ | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 71 | ((val1, type1, fixity1, structure1, signature1, functor1), | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 72 | (val2, type2, fixity2, structure2, signature2, functor2)) : T = | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 73 | (Symtab.merge (K true) (val1, val2), | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 74 | Symtab.merge (K true) (type1, type2), | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 75 | Symtab.merge (K true) (fixity1, fixity2), | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 76 | Symtab.merge (K true) (structure1, structure2), | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 77 | Symtab.merge (K true) (signature1, signature2), | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 78 | Symtab.merge (K true) (functor1, functor2)); | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 79 | ); | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 80 | |
| 28279 | 81 | val inherit_env = Context.proof_map o ML_Env.put o ML_Env.get o Context.Proof; | 
| 82 | ||
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 83 | val name_space: ML_NameSpace.nameSpace = | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 84 | let | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 85 | fun lookup sel1 sel2 name = | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 86 | Context.thread_data () | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 87 | |> (fn NONE => NONE | SOME context => Symtab.lookup (sel1 (ML_Env.get context)) name) | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 88 | |> (fn NONE => sel2 ML_NameSpace.global name | some => some); | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 89 | |
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 90 | fun all sel1 sel2 () = | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 91 | Context.thread_data () | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 92 | |> (fn NONE => [] | SOME context => Symtab.dest (sel1 (ML_Env.get context))) | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 93 | |> append (sel2 ML_NameSpace.global ()) | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 94 | |> sort_distinct (string_ord o pairself #1); | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 95 | |
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 96 | fun enter ap1 sel2 entry = | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 97 | if is_some (Context.thread_data ()) then | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 98 | Context.>> (ML_Env.map (ap1 (Symtab.update entry))) | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 99 | else sel2 ML_NameSpace.global entry; | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 100 | in | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 101 |    {lookupVal    = lookup #1 #lookupVal,
 | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 102 | lookupType = lookup #2 #lookupType, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 103 | lookupFix = lookup #3 #lookupFix, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 104 | lookupStruct = lookup #4 #lookupStruct, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 105 | lookupSig = lookup #5 #lookupSig, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 106 | lookupFunct = lookup #6 #lookupFunct, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 107 | enterVal = enter (fn h => fn (a, b, c, d, e, f) => (h a, b, c, d, e, f)) #enterVal, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 108 | enterType = enter (fn h => fn (a, b, c, d, e, f) => (a, h b, c, d, e, f)) #enterType, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 109 | enterFix = enter (fn h => fn (a, b, c, d, e, f) => (a, b, h c, d, e, f)) #enterFix, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 110 | enterStruct = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, h d, e, f)) #enterStruct, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 111 | enterSig = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, h e, f)) #enterSig, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 112 | enterFunct = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, e, h f)) #enterFunct, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 113 | allVal = all #1 #allVal, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 114 | allType = all #2 #allType, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 115 | allFix = all #3 #allFix, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 116 | allStruct = all #4 #allStruct, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 117 | allSig = all #5 #allSig, | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 118 | allFunct = all #6 #allFunct} | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 119 | end; | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 120 | |
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 121 | |
| 26416 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 122 | (* theorem bindings *) | 
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 123 | |
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 124 | val stored_thms: thm list ref = ref []; | 
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 125 | |
| 26492 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 126 | fun ml_store sel (name, ths) = | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 127 | let | 
| 29579 | 128 | val ths' = Context.>>> (Context.map_theory_result | 
| 129 | (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 | 130 | val _ = | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 131 | if name = "" then () | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 132 | 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 | 133 |         error ("Cannot bind theorem(s) " ^ quote name ^ " as ML value")
 | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 134 | else setmp stored_thms ths' (fn () => | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 135 | use_text name_space (0, "") Output.ml_output true | 
| 26492 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 136 |           ("val " ^ name ^ " = " ^ sel ^ "(! ML_Context.stored_thms);")) ();
 | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 137 | in () end; | 
| 24574 | 138 | |
| 26492 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 139 | val ml_store_thms = ml_store ""; | 
| 
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
 wenzelm parents: 
26473diff
changeset | 140 | 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 | 141 | |
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 142 | fun bind_thm (name, thm) = ml_store_thm (name, Drule.standard thm); | 
| 
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
 wenzelm parents: 
26405diff
changeset | 143 | fun bind_thms (name, thms) = ml_store_thms (name, map Drule.standard thms); | 
| 24574 | 144 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 145 | fun thm name = ProofContext.get_thm (the_local_context ()) name; | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 146 | fun thms name = ProofContext.get_thms (the_local_context ()) name; | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 147 | |
| 24574 | 148 | |
| 149 | ||
| 150 | (** ML antiquotations **) | |
| 151 | ||
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 152 | (* antiquotation commands *) | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 153 | |
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 154 | type antiq = | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 155 |   {struct_name: string, background: Proof.context} ->
 | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 156 | (Proof.context -> string * string) * Proof.context; | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 157 | |
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 158 | local | 
| 24574 | 159 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 160 | val global_parsers = ref (Symtab.empty: | 
| 27868 | 161 | (Position.T -> Context.generic * Args.T list -> antiq * (Context.generic * Args.T list)) | 
| 162 | Symtab.table); | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 163 | |
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 164 | in | 
| 24574 | 165 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 166 | fun add_antiq name scan = CRITICAL (fn () => | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 167 | change global_parsers (fn tab => | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 168 | (if not (Symtab.defined tab name) then () | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 169 |     else warning ("Redefined ML antiquotation: " ^ quote name);
 | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 170 | Symtab.update (name, scan) tab))); | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 171 | |
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 172 | fun antiquotation src ctxt = | 
| 24574 | 173 | let val ((name, _), pos) = Args.dest_src src in | 
| 174 | (case Symtab.lookup (! global_parsers) name of | |
| 175 |       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 | 176 | | SOME scan => (Position.report (Markup.ML_antiq name) pos; | 
| 27895 | 177 | Args.context_syntax "ML antiquotation" (scan pos) src ctxt)) | 
| 24574 | 178 | end; | 
| 179 | ||
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 180 | end; | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 181 | |
| 24574 | 182 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 183 | (* parsing and evaluation *) | 
| 24574 | 184 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 185 | local | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 186 | |
| 24574 | 187 | structure P = OuterParse; | 
| 188 | ||
| 189 | val antiq = | |
| 27809 
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27781diff
changeset | 190 | P.!!! (P.position P.xname -- Args.parse --| Scan.ahead P.eof) | 
| 24574 | 191 | >> (fn ((x, pos), y) => Args.src ((x, y), pos)); | 
| 192 | ||
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 193 | val struct_name = "Isabelle"; | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 194 | |
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 195 | fun eval_antiquotes (txt, pos) opt_context = | 
| 24574 | 196 | let | 
| 27874 
f0364f1c0ecf
antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
 wenzelm parents: 
27868diff
changeset | 197 | val syms = SymbolPos.explode (txt, pos); | 
| 
f0364f1c0ecf
antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
 wenzelm parents: 
27868diff
changeset | 198 | val ants = Antiquote.read (syms, pos); | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 199 | 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 | 200 | val ((ml_env, ml_body), opt_ctxt') = | 
| 27874 
f0364f1c0ecf
antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
 wenzelm parents: 
27868diff
changeset | 201 | if not (exists Antiquote.is_antiq ants) | 
| 
f0364f1c0ecf
antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
 wenzelm parents: 
27868diff
changeset | 202 |       then (("", Symbol.escape (SymbolPos.content syms)), opt_ctxt)
 | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 203 | else | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 204 | let | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 205 | val ctxt = | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 206 | (case opt_ctxt of | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 207 |               NONE => error ("Unknown context -- cannot expand ML antiquotations" ^
 | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 208 | Position.str_of pos) | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 209 | | SOME ctxt => Context.proof_of ctxt); | 
| 27378 | 210 | |
| 27359 | 211 | val lex = #1 (OuterKeyword.get_lexicons ()); | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 212 |           fun no_decl _ = ("", "");
 | 
| 27378 | 213 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 214 |           fun expand (Antiquote.Text s) state = (K ("", Symbol.escape s), state)
 | 
| 27781 | 215 | | expand (Antiquote.Antiq x) (scope, background) = | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 216 | let | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 217 | val context = Stack.top scope; | 
| 27781 | 218 | val (f, context') = antiquotation (Antiquote.read_antiq lex antiq x) context; | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 219 |                   val (decl, background') = f {background = background, struct_name = struct_name};
 | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 220 | in (decl, (Stack.map_top (K context') scope, background')) end | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 221 | | expand (Antiquote.Open _) (scope, background) = | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 222 | (no_decl, (Stack.push scope, background)) | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 223 | | expand (Antiquote.Close _) (scope, background) = | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 224 | (no_decl, (Stack.pop scope, background)); | 
| 24574 | 225 | |
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 226 | val (decls, (_, ctxt')) = fold_map expand ants (Stack.init ctxt, ctxt); | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 227 | val ml = decls |> map (fn decl => decl ctxt') |> split_list |> pairself implode; | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 228 | in (ml, SOME (Context.Proof ctxt')) end; | 
| 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 229 |   in (("structure " ^ struct_name ^ " =\nstruct\n" ^ ml_env ^ "end;", ml_body), opt_ctxt') end;
 | 
| 24574 | 230 | |
| 231 | in | |
| 232 | ||
| 233 | val trace = ref false; | |
| 234 | ||
| 26385 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 wenzelm parents: 
26374diff
changeset | 235 | fun eval_wrapper pr verbose pos txt = | 
| 24574 | 236 | let | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 237 | fun eval_raw p = use_text name_space | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 238 | (the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p)) pr; | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 239 | |
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 240 | (*prepare source text*) | 
| 27878 | 241 | val _ = Position.report Markup.ML_source pos; | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 242 | val ((env, body), env_ctxt) = eval_antiquotes (txt, pos) (Context.thread_data ()); | 
| 27343 
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
 wenzelm parents: 
26881diff
changeset | 243 | val _ = if ! trace then tracing (cat_lines [env, body]) else (); | 
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 244 | |
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 245 | (*prepare static ML environment*) | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 246 | val _ = Context.setmp_thread_data env_ctxt | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 247 | (fn () => (eval_raw Position.none false env; Context.thread_data ())) () | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 248 | |> (fn NONE => () | SOME context' => Context.>> (ML_Env.put (ML_Env.get context'))); | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 249 | |
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 250 | (*eval ML*) | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 251 | val _ = eval_raw pos verbose body; | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 252 | |
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 253 | (*reset static ML environment*) | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 254 | val _ = eval_raw Position.none false "structure Isabelle = struct end"; | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 255 | in () end; | 
| 24574 | 256 | |
| 257 | end; | |
| 258 | ||
| 259 | ||
| 260 | (* ML evaluation *) | |
| 261 | ||
| 26455 | 262 | val eval = eval_wrapper Output.ml_output; | 
| 26881 | 263 | fun eval_file path = eval true (Path.position path) (File.read path); | 
| 24574 | 264 | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 265 | fun eval_in ctxt verbose pos txt = | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 266 | Context.setmp_thread_data (Option.map Context.Proof ctxt) (fn () => eval verbose pos txt) (); | 
| 25204 | 267 | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 268 | fun evaluate ctxt pr verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () => | 
| 25204 | 269 | let | 
| 270 | val _ = r := NONE; | |
| 28270 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 271 | val _ = Context.setmp_thread_data (SOME (Context.Proof ctxt)) (fn () => | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 272 | eval_wrapper pr verbose Position.none | 
| 
7ada24ebea94
explicit handling of ML environment within generic context;
 wenzelm parents: 
27895diff
changeset | 273 |         ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))")) ();
 | 
| 25204 | 274 |   in (case ! r of NONE => error ("Bad evaluation for " ^ ref_name) | SOME e => e) end) ();
 | 
| 24574 | 275 | |
| 26455 | 276 | fun expression pos bind body txt = | 
| 277 | exec (fn () => eval false pos | |
| 278 |     ("Context.set_thread_data (SOME (let " ^ bind ^ " = " ^ txt ^ " in " ^ body ^
 | |
| 279 | " end (ML_Context.the_generic_context ())));")); | |
| 280 | ||
| 24574 | 281 | end; | 
| 282 | ||
| 283 | structure Basic_ML_Context: BASIC_ML_CONTEXT = ML_Context; | |
| 284 | open Basic_ML_Context; |