author | wenzelm |
Sat, 28 Jun 2008 15:17:26 +0200 | |
changeset 27378 | 0968c0d0b969 |
parent 27359 | 54b5367a827a |
child 27723 | ce8f79b91ed1 |
permissions | -rw-r--r-- |
24581 | 1 |
(* Title: Pure/ML/ml_context.ML |
24574 | 2 |
ID: $Id$ |
3 |
Author: Makarius |
|
4 |
||
5 |
ML context and antiquotations. |
|
6 |
*) |
|
7 |
||
8 |
signature BASIC_ML_CONTEXT = |
|
9 |
sig |
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
10 |
val bind_thm: string * thm -> unit |
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
11 |
val bind_thms: string * thm list -> unit |
24574 | 12 |
val thm: xstring -> thm |
13 |
val thms: xstring -> thm list |
|
25204 | 14 |
end |
24574 | 15 |
|
16 |
signature ML_CONTEXT = |
|
17 |
sig |
|
18 |
include BASIC_ML_CONTEXT |
|
19 |
val the_generic_context: unit -> Context.generic |
|
26432
095e448b95a0
renamed ML_Context.the_context to ML_Context.the_global_context;
wenzelm
parents:
26416
diff
changeset
|
20 |
val the_global_context: unit -> theory |
24574 | 21 |
val the_local_context: unit -> Proof.context |
26455 | 22 |
val exec: (unit -> unit) -> Context.generic -> Context.generic |
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
23 |
val stored_thms: thm list ref |
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
24 |
val ml_store_thm: string * thm -> unit |
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
25 |
val ml_store_thms: string * thm list -> unit |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
26 |
type antiq = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
27 |
{struct_name: string, background: Proof.context} -> |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
28 |
(Proof.context -> string * string) * Proof.context |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
29 |
val add_antiq: string -> |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
30 |
(Context.generic * Args.T list -> antiq * (Context.generic * Args.T list)) -> unit |
24574 | 31 |
val trace: bool ref |
26455 | 32 |
val eval: bool -> Position.T -> string -> unit |
33 |
val eval_file: Path.T -> unit |
|
34 |
val eval_in: Context.generic option -> bool -> Position.T -> string -> unit |
|
25204 | 35 |
val evaluate: (string -> unit) * (string -> 'b) -> bool -> |
36 |
string * (unit -> 'a) option ref -> string -> 'a |
|
26455 | 37 |
val expression: Position.T -> string -> string -> string -> Context.generic -> Context.generic |
25204 | 38 |
end |
24574 | 39 |
|
40 |
structure ML_Context: ML_CONTEXT = |
|
41 |
struct |
|
42 |
||
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
43 |
(** implicit ML context **) |
24574 | 44 |
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
45 |
val the_generic_context = Context.the_thread_data; |
26432
095e448b95a0
renamed ML_Context.the_context to ML_Context.the_global_context;
wenzelm
parents:
26416
diff
changeset
|
46 |
val the_global_context = Context.theory_of o the_generic_context; |
24574 | 47 |
val the_local_context = Context.proof_of o the_generic_context; |
48 |
||
26455 | 49 |
fun exec (e: unit -> unit) context = |
50 |
(case Context.setmp_thread_data (SOME context) (fn () => (e (); Context.thread_data ())) () of |
|
51 |
SOME context' => context' |
|
52 |
| NONE => error "Missing context after execution"); |
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
53 |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
54 |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
55 |
(* theorem bindings *) |
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
56 |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
57 |
val stored_thms: thm list ref = ref []; |
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
58 |
|
26492
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
59 |
fun ml_store sel (name, ths) = |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
60 |
let |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
61 |
val ths' = Context.>>> (Context.map_theory_result (PureThy.store_thms (name, ths))); |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
62 |
val _ = |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
63 |
if name = "" then () |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
64 |
else if not (ML_Syntax.is_identifier name) then |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
65 |
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:
26473
diff
changeset
|
66 |
else setmp stored_thms ths' (fn () => |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
67 |
use_text (0, "") Output.ml_output true |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
68 |
("val " ^ name ^ " = " ^ sel ^ "(! ML_Context.stored_thms);")) (); |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
69 |
in () end; |
24574 | 70 |
|
26492
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
71 |
val ml_store_thms = ml_store ""; |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
72 |
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:
26405
diff
changeset
|
73 |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
74 |
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:
26405
diff
changeset
|
75 |
fun bind_thms (name, thms) = ml_store_thms (name, map Drule.standard thms); |
24574 | 76 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
77 |
fun thm name = ProofContext.get_thm (the_local_context ()) name; |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
78 |
fun thms name = ProofContext.get_thms (the_local_context ()) name; |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
79 |
|
24574 | 80 |
|
81 |
||
82 |
(** ML antiquotations **) |
|
83 |
||
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
84 |
(* antiquotation commands *) |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
85 |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
86 |
type antiq = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
87 |
{struct_name: string, background: Proof.context} -> |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
88 |
(Proof.context -> string * string) * Proof.context; |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
89 |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
90 |
local |
24574 | 91 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
92 |
val global_parsers = ref (Symtab.empty: |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
93 |
(Context.generic * Args.T list -> antiq * (Context.generic * Args.T list)) Symtab.table); |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
94 |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
95 |
in |
24574 | 96 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
97 |
fun add_antiq name scan = CRITICAL (fn () => |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
98 |
change global_parsers (fn tab => |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
99 |
(if not (Symtab.defined tab name) then () |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
100 |
else warning ("Redefined ML antiquotation: " ^ quote name); |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
101 |
Symtab.update (name, scan) tab))); |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
102 |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
103 |
fun antiquotation src ctxt = |
24574 | 104 |
let val ((name, _), pos) = Args.dest_src src in |
105 |
(case Symtab.lookup (! global_parsers) name of |
|
106 |
NONE => error ("Unknown ML antiquotation command: " ^ quote name ^ Position.str_of pos) |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
107 |
| SOME scan => Args.context_syntax "ML antiquotation" scan src ctxt) |
24574 | 108 |
end; |
109 |
||
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
110 |
end; |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
111 |
|
24574 | 112 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
113 |
(* parsing and evaluation *) |
24574 | 114 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
115 |
local |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
116 |
|
24574 | 117 |
structure P = OuterParse; |
118 |
||
119 |
val antiq = |
|
120 |
P.!!! (P.position P.xname -- P.arguments --| Scan.ahead P.eof) |
|
121 |
>> (fn ((x, pos), y) => Args.src ((x, y), pos)); |
|
122 |
||
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
123 |
fun eval_antiquotes struct_name (txt, pos) opt_ctxt = |
24574 | 124 |
let |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
125 |
val ants = Antiquote.scan_antiquotes (txt, pos); |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
126 |
val ((ml_env, ml_body), opt_ctxt') = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
127 |
if not (exists Antiquote.is_antiq ants) then (("", Symbol.escape txt), opt_ctxt) |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
128 |
else |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
129 |
let |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
130 |
val ctxt = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
131 |
(case opt_ctxt of |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
132 |
NONE => error ("Unknown context -- cannot expand ML antiquotations" ^ |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
133 |
Position.str_of pos) |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
134 |
| SOME context => Context.proof_of context); |
27378 | 135 |
|
27359 | 136 |
val lex = #1 (OuterKeyword.get_lexicons ()); |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
137 |
fun no_decl _ = ("", ""); |
27378 | 138 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
139 |
fun expand (Antiquote.Text s) state = (K ("", Symbol.escape s), state) |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
140 |
| expand (Antiquote.Antiq x) (scope, background) = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
141 |
let |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
142 |
val context = Stack.top scope; |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
143 |
val (f, context') = antiquotation (Antiquote.scan_arguments lex antiq x) context; |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
144 |
val (decl, background') = f {background = background, struct_name = struct_name}; |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
145 |
in (decl, (Stack.map_top (K context') scope, background')) end |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
146 |
| expand (Antiquote.Open _) (scope, background) = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
147 |
(no_decl, (Stack.push scope, background)) |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
148 |
| expand (Antiquote.Close _) (scope, background) = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
149 |
(no_decl, (Stack.pop scope, background)); |
24574 | 150 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
151 |
val (decls, (_, ctxt')) = fold_map expand ants (Stack.init ctxt, ctxt); |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
152 |
val ml = decls |> map (fn decl => decl ctxt') |> split_list |> pairself implode; |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
153 |
in (ml, SOME (Context.Proof ctxt')) end; |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
154 |
in (("structure " ^ struct_name ^ " =\nstruct\n" ^ ml_env ^ "end;", ml_body), opt_ctxt') end; |
24574 | 155 |
|
156 |
in |
|
157 |
||
158 |
val trace = ref false; |
|
159 |
||
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
160 |
fun eval_wrapper pr verbose pos txt = |
24574 | 161 |
let |
26473
2266e5fd7b63
eval_wrapper: non-critical version via unique Isabelle structure, proper forget_structure;
wenzelm
parents:
26455
diff
changeset
|
162 |
val struct_name = |
2266e5fd7b63
eval_wrapper: non-critical version via unique Isabelle structure, proper forget_structure;
wenzelm
parents:
26455
diff
changeset
|
163 |
if Multithreading.available then "Isabelle" ^ serial_string () |
2266e5fd7b63
eval_wrapper: non-critical version via unique Isabelle structure, proper forget_structure;
wenzelm
parents:
26455
diff
changeset
|
164 |
else "Isabelle"; |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
165 |
val ((env, body), env_ctxt) = eval_antiquotes struct_name (txt, pos) (Context.thread_data ()); |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
166 |
val _ = if ! trace then tracing (cat_lines [env, body]) else (); |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
167 |
fun eval_raw p = |
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
168 |
use_text (the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p)) pr; |
25755 | 169 |
in |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
170 |
Context.setmp_thread_data env_ctxt (fn () => eval_raw Position.none false env) (); |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
171 |
NAMED_CRITICAL "ML" (fn () => eval_raw pos verbose body); (* FIXME non-critical with local ML env *) |
26473
2266e5fd7b63
eval_wrapper: non-critical version via unique Isabelle structure, proper forget_structure;
wenzelm
parents:
26455
diff
changeset
|
172 |
forget_structure struct_name |
25755 | 173 |
end; |
24574 | 174 |
|
175 |
end; |
|
176 |
||
177 |
||
178 |
(* ML evaluation *) |
|
179 |
||
26455 | 180 |
val eval = eval_wrapper Output.ml_output; |
26881 | 181 |
fun eval_file path = eval true (Path.position path) (File.read path); |
24574 | 182 |
|
26455 | 183 |
fun eval_in context verbose pos txt = |
184 |
Context.setmp_thread_data context (fn () => eval verbose pos txt) (); |
|
25204 | 185 |
|
25700 | 186 |
fun evaluate pr verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () => |
25204 | 187 |
let |
188 |
val _ = r := NONE; |
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
189 |
val _ = eval_wrapper pr verbose Position.none |
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
190 |
("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))"); |
25204 | 191 |
in (case ! r of NONE => error ("Bad evaluation for " ^ ref_name) | SOME e => e) end) (); |
24574 | 192 |
|
26455 | 193 |
fun expression pos bind body txt = |
194 |
exec (fn () => eval false pos |
|
195 |
("Context.set_thread_data (SOME (let " ^ bind ^ " = " ^ txt ^ " in " ^ body ^ |
|
196 |
" end (ML_Context.the_generic_context ())));")); |
|
197 |
||
24574 | 198 |
end; |
199 |
||
200 |
structure Basic_ML_Context: BASIC_ML_CONTEXT = ML_Context; |
|
201 |
open Basic_ML_Context; |
|
27378 | 202 |