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