| author | wenzelm |
| Wed, 26 Mar 2008 22:39:57 +0100 | |
| changeset 26409 | 1ceabad5a2c8 |
| parent 26405 | 0cb6f2013e99 |
| child 26416 | a66f86ef7bb9 |
| 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 |
|
10 |
val the_context: unit -> theory |
|
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 get_context: unit -> Context.generic option |
|
19 |
val set_context: Context.generic option -> unit |
|
20 |
val setmp: Context.generic option -> ('a -> 'b) -> 'a -> 'b
|
|
21 |
val the_generic_context: unit -> Context.generic |
|
22 |
val the_local_context: unit -> Proof.context |
|
23 |
val pass_context: Context.generic -> ('a -> 'b) -> 'a -> 'b * Context.generic
|
|
24 |
val >> : (theory -> theory) -> unit |
|
25 |
val add_keywords: string list -> unit |
|
26 |
val inline_antiq: string -> |
|
27 |
(Context.generic * Args.T list -> string * (Context.generic * Args.T list)) -> unit |
|
28 |
val value_antiq: string -> |
|
29 |
(Context.generic * Args.T list -> (string * string) * (Context.generic * Args.T list)) -> unit |
|
30 |
val proj_value_antiq: string -> (Context.generic * Args.T list -> |
|
31 |
(string * string * string) * (Context.generic * Args.T list)) -> unit |
|
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
32 |
val eval_antiquotes_fn: (string * Position.T -> string * string) ref (* FIXME tmp *) |
| 24574 | 33 |
val trace: bool ref |
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
34 |
val use_mltext: bool -> Position.T -> string -> Context.generic option -> unit |
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
35 |
val use_mltext_update: bool -> Position.T -> string -> Context.generic -> Context.generic |
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
36 |
val use_let: Position.T -> string -> string -> string -> Context.generic -> Context.generic |
| 24574 | 37 |
val use: Path.T -> unit |
| 25204 | 38 |
val evaluate: (string -> unit) * (string -> 'b) -> bool -> |
39 |
string * (unit -> 'a) option ref -> string -> 'a |
|
40 |
end |
|
| 24574 | 41 |
|
42 |
structure ML_Context: ML_CONTEXT = |
|
43 |
struct |
|
44 |
||
45 |
(** Implicit ML context **) |
|
46 |
||
47 |
local |
|
48 |
val current_context = ref (NONE: Context.generic option); |
|
49 |
in |
|
50 |
fun get_context () = ! current_context; |
|
51 |
fun set_context opt_context = current_context := opt_context; |
|
52 |
fun setmp opt_context f x = Library.setmp current_context opt_context f x; |
|
53 |
end; |
|
54 |
||
55 |
fun the_generic_context () = |
|
56 |
(case get_context () of |
|
57 |
SOME context => context |
|
58 |
| _ => error "Unknown context"); |
|
59 |
||
60 |
val the_local_context = Context.proof_of o the_generic_context; |
|
61 |
||
62 |
val the_context = Context.theory_of o the_generic_context; |
|
63 |
||
64 |
fun pass_context context f x = |
|
| 26405 | 65 |
(case setmp (SOME context) (fn () => (f x, get_context ())) () of |
| 24574 | 66 |
(y, SOME context') => (y, context') |
67 |
| (_, NONE) => error "Lost context in ML"); |
|
68 |
||
| 25700 | 69 |
fun change_theory f = NAMED_CRITICAL "ML" (fn () => |
| 24574 | 70 |
set_context (SOME (Context.map_theory f (the_generic_context ())))); |
71 |
||
72 |
||
73 |
||
74 |
(** ML antiquotations **) |
|
75 |
||
76 |
(* maintain keywords *) |
|
77 |
||
78 |
val global_lexicon = ref Scan.empty_lexicon; |
|
79 |
||
80 |
fun add_keywords keywords = CRITICAL (fn () => |
|
81 |
change global_lexicon (Scan.extend_lexicon (map Symbol.explode keywords))); |
|
82 |
||
83 |
||
84 |
(* maintain commands *) |
|
85 |
||
86 |
datatype antiq = Inline of string | ProjValue of string * string * string; |
|
87 |
val global_parsers = ref (Symtab.empty: |
|
88 |
(Context.generic * Args.T list -> antiq * (Context.generic * Args.T list)) Symtab.table); |
|
89 |
||
90 |
local |
|
91 |
||
92 |
fun add_item kind name scan = CRITICAL (fn () => |
|
93 |
change global_parsers (fn tab => |
|
94 |
(if not (Symtab.defined tab name) then () |
|
95 |
else warning ("Redefined ML antiquotation: " ^ quote name);
|
|
96 |
Symtab.update (name, scan >> kind) tab))); |
|
97 |
||
98 |
in |
|
99 |
||
100 |
val inline_antiq = add_item Inline; |
|
101 |
val proj_value_antiq = add_item ProjValue; |
|
102 |
fun value_antiq name scan = proj_value_antiq name (scan >> (fn (a, s) => (a, s, ""))); |
|
103 |
||
104 |
end; |
|
105 |
||
106 |
||
107 |
(* command syntax *) |
|
108 |
||
109 |
fun syntax scan src = |
|
110 |
#1 (Args.context_syntax "ML antiquotation" scan src (the_local_context ())); |
|
111 |
||
112 |
fun command src = |
|
113 |
let val ((name, _), pos) = Args.dest_src src in |
|
114 |
(case Symtab.lookup (! global_parsers) name of |
|
115 |
NONE => error ("Unknown ML antiquotation command: " ^ quote name ^ Position.str_of pos)
|
|
116 |
| SOME scan => syntax scan src) |
|
117 |
end; |
|
118 |
||
119 |
||
120 |
(* outer syntax *) |
|
121 |
||
122 |
structure T = OuterLex; |
|
123 |
structure P = OuterParse; |
|
124 |
||
125 |
val antiq = |
|
126 |
P.!!! (P.position P.xname -- P.arguments --| Scan.ahead P.eof) |
|
127 |
>> (fn ((x, pos), y) => Args.src ((x, y), pos)); |
|
128 |
||
129 |
||
130 |
(* input/output wrappers *) |
|
131 |
||
132 |
local |
|
133 |
||
134 |
val isabelle_struct = enclose "structure Isabelle =\nstruct\n" "end;"; |
|
135 |
val isabelle_struct0 = isabelle_struct ""; |
|
136 |
||
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
137 |
fun eval_antiquotes txt_pos = |
| 24574 | 138 |
let |
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
139 |
val ants = Antiquote.scan_antiquotes txt_pos; |
| 24574 | 140 |
|
141 |
fun expand (Antiquote.Text s) names = (("", Symbol.escape s), names)
|
|
142 |
| expand (Antiquote.Antiq x) names = |
|
143 |
(case command (Antiquote.scan_arguments (! global_lexicon) antiq x) of |
|
144 |
Inline s => (("", s), names)
|
|
145 |
| ProjValue (a, s, f) => |
|
146 |
let |
|
147 |
val a' = if ML_Syntax.is_identifier a then a else "val"; |
|
148 |
val ([b], names') = Name.variants [a'] names; |
|
149 |
val binding = "val " ^ b ^ " = " ^ s ^ ";\n"; |
|
150 |
val value = |
|
151 |
if f = "" then "Isabelle." ^ b |
|
152 |
else "(" ^ f ^ " Isabelle." ^ b ^ ")";
|
|
153 |
in ((binding, value), names') end); |
|
154 |
||
155 |
val (decls, body) = |
|
| 25700 | 156 |
NAMED_CRITICAL "ML" (fn () => fold_map expand ants ML_Syntax.reserved) |
| 24574 | 157 |
|> #1 |> split_list |> pairself implode; |
| 25700 | 158 |
in (isabelle_struct decls, body) end; |
| 24574 | 159 |
|
160 |
in |
|
161 |
||
162 |
val eval_antiquotes_fn = ref eval_antiquotes; |
|
163 |
||
164 |
val trace = ref false; |
|
165 |
||
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
166 |
fun eval_wrapper pr verbose pos txt = |
| 24574 | 167 |
let |
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
168 |
val (txt1, txt2) = ! eval_antiquotes_fn (txt, pos); (* FIXME tmp hook *) |
| 24574 | 169 |
val _ = if ! trace then tracing (cat_lines [txt1, txt2]) else (); |
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
170 |
fun eval p = |
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
171 |
use_text (the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p)) pr; |
| 25755 | 172 |
in |
173 |
NAMED_CRITICAL "ML" (fn () => |
|
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
174 |
(eval Position.none false txt1; |
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
175 |
eval pos verbose txt2; |
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
176 |
eval Position.none false isabelle_struct0)) |
| 25755 | 177 |
end; |
| 24574 | 178 |
|
179 |
end; |
|
180 |
||
181 |
||
182 |
(* ML evaluation *) |
|
183 |
||
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
184 |
fun use_mltext verbose pos txt opt_context = |
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
185 |
setmp opt_context (fn () => eval_wrapper Output.ml_output verbose pos txt) (); |
| 24574 | 186 |
|
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
187 |
fun use_mltext_update verbose pos txt context = |
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
188 |
#2 (pass_context context (fn () => eval_wrapper Output.ml_output verbose pos txt) ()); |
| 24574 | 189 |
|
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
190 |
fun use_let pos bind body txt = |
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
191 |
use_mltext_update false pos |
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
192 |
("ML_Context.set_context (SOME (let " ^ bind ^ " = " ^ txt ^ " in " ^ body ^
|
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
193 |
" end (ML_Context.the_generic_context ())));"); |
| 24574 | 194 |
|
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
195 |
fun use path = eval_wrapper Output.ml_output true (Position.path path) (File.read path); |
| 25204 | 196 |
|
| 25700 | 197 |
fun evaluate pr verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () => |
| 25204 | 198 |
let |
199 |
val _ = r := NONE; |
|
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
200 |
val _ = eval_wrapper pr verbose Position.none |
|
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
201 |
("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))");
|
| 25204 | 202 |
in (case ! r of NONE => error ("Bad evaluation for " ^ ref_name) | SOME e => e) end) ();
|
| 24574 | 203 |
|
204 |
||
205 |
(* basic antiquotations *) |
|
206 |
||
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
207 |
fun context x = (Scan.state >> Context.proof_of) x; |
| 25142 | 208 |
|
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
209 |
local |
| 25142 | 210 |
|
| 24574 | 211 |
val _ = value_antiq "context" (Scan.succeed ("context", "ML_Context.the_local_context ()"));
|
212 |
||
| 25142 | 213 |
val _ = inline_antiq "sort" (context -- Scan.lift Args.name >> (fn (ctxt, s) => |
214 |
ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s)))); |
|
215 |
||
| 24574 | 216 |
val _ = inline_antiq "typ" (Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ)); |
217 |
val _ = inline_antiq "term" (Args.term >> (ML_Syntax.atomic o ML_Syntax.print_term)); |
|
218 |
val _ = inline_antiq "prop" (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term)); |
|
219 |
||
220 |
val _ = value_antiq "ctyp" (Args.typ >> (fn T => |
|
221 |
("ctyp",
|
|
222 |
"Thm.ctyp_of (ML_Context.the_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_typ T)))); |
|
223 |
||
224 |
val _ = value_antiq "cterm" (Args.term >> (fn t => |
|
225 |
("cterm",
|
|
226 |
"Thm.cterm_of (ML_Context.the_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t)))); |
|
227 |
||
228 |
val _ = value_antiq "cprop" (Args.prop >> (fn t => |
|
229 |
("cprop",
|
|
230 |
"Thm.cterm_of (ML_Context.the_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t)))); |
|
231 |
||
| 25142 | 232 |
fun type_ syn = (context -- Scan.lift Args.name >> (fn (ctxt, c) => |
| 24695 | 233 |
#1 (Term.dest_Type (ProofContext.read_tyname ctxt c)) |
| 25142 | 234 |
|> syn ? Sign.base_name |
| 24695 | 235 |
|> ML_Syntax.print_string)); |
| 24574 | 236 |
|
| 25142 | 237 |
val _ = inline_antiq "type_name" (type_ false); |
238 |
val _ = inline_antiq "type_syntax" (type_ true); |
|
239 |
||
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
240 |
fun const syn = context -- Scan.lift Args.name >> (fn (ctxt, c) => |
|
25346
7f2e3292e3dd
renamed ProofContext.read_const' to ProofContext.read_const_proper;
wenzelm
parents:
25334
diff
changeset
|
241 |
#1 (Term.dest_Const (ProofContext.read_const_proper ctxt c)) |
| 24574 | 242 |
|> syn ? ProofContext.const_syntax_name ctxt |
243 |
|> ML_Syntax.print_string); |
|
244 |
||
245 |
val _ = inline_antiq "const_name" (const false); |
|
246 |
val _ = inline_antiq "const_syntax" (const true); |
|
247 |
||
| 26188 | 248 |
val _ = inline_antiq "const" |
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
249 |
(context -- Scan.lift Args.name -- |
| 26188 | 250 |
Scan.optional (Scan.lift (Args.$$$ "(") |-- Args.enum1 "," Args.typ --| Scan.lift (Args.$$$ ")")) []
|
251 |
>> (fn ((ctxt, c), Ts) => |
|
252 |
let |
|
253 |
val (c, _) = Term.dest_Const (ProofContext.read_const_proper ctxt c); |
|
| 26268 | 254 |
in ML_Syntax.atomic (ML_Syntax.print_term (ProofContext.mk_const ctxt (c, Ts))) end)); |
| 26188 | 255 |
|
| 25142 | 256 |
in val _ = () end; |
| 24574 | 257 |
|
258 |
||
259 |
(** fact references **) |
|
260 |
||
|
26346
17debd2fff8e
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26336
diff
changeset
|
261 |
fun thm name = ProofContext.get_thm (the_local_context ()) name; |
|
17debd2fff8e
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26336
diff
changeset
|
262 |
fun thms name = ProofContext.get_thms (the_local_context ()) name; |
| 24574 | 263 |
|
264 |
||
265 |
local |
|
266 |
||
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26268
diff
changeset
|
267 |
fun print_interval (Facts.FromTo arg) = |
|
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26268
diff
changeset
|
268 |
"Facts.FromTo " ^ ML_Syntax.print_pair ML_Syntax.print_int ML_Syntax.print_int arg |
|
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26268
diff
changeset
|
269 |
| print_interval (Facts.From i) = "Facts.From " ^ ML_Syntax.print_int i |
|
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26268
diff
changeset
|
270 |
| print_interval (Facts.Single i) = "Facts.Single " ^ ML_Syntax.print_int i; |
| 24574 | 271 |
|
|
26368
3e74b09ce466
thm_antiq: produce error at runtime, not compile time;
wenzelm
parents:
26361
diff
changeset
|
272 |
fun thm_antiq kind get get_name = value_antiq kind |
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26374
diff
changeset
|
273 |
(context -- Scan.lift (Args.position Args.name -- Scan.option Args.thm_sel) |
| 26374 | 274 |
>> (fn (ctxt, ((name, pos), sel)) => |
|
26368
3e74b09ce466
thm_antiq: produce error at runtime, not compile time;
wenzelm
parents:
26361
diff
changeset
|
275 |
let |
| 26374 | 276 |
val _ = get ctxt (Facts.Named ((name, pos), sel)); |
277 |
val txt = |
|
278 |
"(Facts.Named ((" ^ ML_Syntax.print_string name ^ ", Position.none), " ^
|
|
279 |
ML_Syntax.print_option (ML_Syntax.print_list print_interval) sel ^ "))"; |
|
280 |
in (name, get_name ^ " (ML_Context.the_local_context ()) " ^ txt) end)); |
|
| 24574 | 281 |
|
282 |
in |
|
283 |
||
284 |
val _ = add_keywords ["(", ")", "-", ","];
|
|
|
26368
3e74b09ce466
thm_antiq: produce error at runtime, not compile time;
wenzelm
parents:
26361
diff
changeset
|
285 |
val _ = thm_antiq "thm" ProofContext.get_fact_single "ProofContext.get_fact_single"; |
|
3e74b09ce466
thm_antiq: produce error at runtime, not compile time;
wenzelm
parents:
26361
diff
changeset
|
286 |
val _ = thm_antiq "thms" ProofContext.get_fact "ProofContext.get_fact"; |
| 24574 | 287 |
|
288 |
end; |
|
289 |
||
| 24685 | 290 |
val _ = proj_value_antiq "cpat" (Scan.lift Args.name >> |
291 |
(fn name => ("cpat",
|
|
292 |
"Thm.cterm_of (ML_Context.the_context ()) (Syntax.read_term \ |
|
293 |
\(ProofContext.set_mode ProofContext.mode_pattern (ML_Context.the_local_context ()))" |
|
294 |
^ ML_Syntax.print_string name ^ ")", ""))); |
|
| 24574 | 295 |
|
296 |
(*final declarations of this structure!*) |
|
297 |
nonfix >>; |
|
298 |
fun >> f = change_theory f; |
|
299 |
||
300 |
end; |
|
301 |
||
302 |
structure Basic_ML_Context: BASIC_ML_CONTEXT = ML_Context; |
|
303 |
open Basic_ML_Context; |