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 |
|
|
74 |
fun save f x = CRITICAL (fn () => setmp (get_context ()) f x);
|
|
75 |
|
|
76 |
fun change_theory f = CRITICAL (fn () =>
|
|
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 |
|
|
144 |
fun eval_antiquotes txt = CRITICAL (fn () =>
|
|
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) =
|
|
163 |
fold_map expand ants ML_Syntax.reserved
|
|
164 |
|> #1 |> split_list |> pairself implode;
|
|
165 |
in (isabelle_struct decls, body) end);
|
|
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;
|
24574
|
178 |
in eval "" false txt1; eval name verbose txt2; eval "" false isabelle_struct0 end;
|
|
179 |
|
25204
|
180 |
fun ML_wrapper pr = eval_wrapper "ML" pr;
|
|
181 |
|
24574
|
182 |
end;
|
|
183 |
|
|
184 |
|
|
185 |
(* ML evaluation *)
|
|
186 |
|
|
187 |
fun use_mltext txt verbose opt_context =
|
25204
|
188 |
setmp opt_context (fn () => ML_wrapper Output.ml_output verbose txt) ();
|
24574
|
189 |
|
|
190 |
fun use_mltext_update txt verbose context =
|
25204
|
191 |
#2 (pass_context context (ML_wrapper Output.ml_output verbose) txt);
|
24574
|
192 |
|
|
193 |
fun use_context txt = use_mltext_update
|
|
194 |
("ML_Context.set_context (SOME ((" ^ txt ^ ") (ML_Context.the_generic_context ())));") false;
|
|
195 |
|
|
196 |
fun use_let bind body txt =
|
|
197 |
use_context ("let " ^ bind ^ " = " ^ txt ^ " in\n" ^ body ^ " end");
|
|
198 |
|
25204
|
199 |
fun use path = eval_wrapper (Path.implode path) Output.ml_output true (File.read path);
|
|
200 |
|
|
201 |
fun evaluate pr verbose (ref_name, r) txt = CRITICAL (fn () =>
|
|
202 |
let
|
|
203 |
val _ = r := NONE;
|
|
204 |
val _ = ML_wrapper pr verbose ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))");
|
|
205 |
in (case ! r of NONE => error ("Bad evaluation for " ^ ref_name) | SOME e => e) end) ();
|
24574
|
206 |
|
|
207 |
|
|
208 |
(* basic antiquotations *)
|
|
209 |
|
25142
|
210 |
local
|
|
211 |
|
|
212 |
fun context x = (Scan.state >> Context.proof_of) x;
|
|
213 |
|
24574
|
214 |
val _ = value_antiq "context" (Scan.succeed ("context", "ML_Context.the_local_context ()"));
|
|
215 |
|
25142
|
216 |
val _ = inline_antiq "sort" (context -- Scan.lift Args.name >> (fn (ctxt, s) =>
|
|
217 |
ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s))));
|
|
218 |
|
24574
|
219 |
val _ = inline_antiq "typ" (Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ));
|
|
220 |
val _ = inline_antiq "term" (Args.term >> (ML_Syntax.atomic o ML_Syntax.print_term));
|
|
221 |
val _ = inline_antiq "prop" (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term));
|
|
222 |
|
|
223 |
val _ = value_antiq "ctyp" (Args.typ >> (fn T =>
|
|
224 |
("ctyp",
|
|
225 |
"Thm.ctyp_of (ML_Context.the_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_typ T))));
|
|
226 |
|
|
227 |
val _ = value_antiq "cterm" (Args.term >> (fn t =>
|
|
228 |
("cterm",
|
|
229 |
"Thm.cterm_of (ML_Context.the_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t))));
|
|
230 |
|
|
231 |
val _ = value_antiq "cprop" (Args.prop >> (fn t =>
|
|
232 |
("cprop",
|
|
233 |
"Thm.cterm_of (ML_Context.the_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t))));
|
|
234 |
|
25142
|
235 |
fun type_ syn = (context -- Scan.lift Args.name >> (fn (ctxt, c) =>
|
24695
|
236 |
#1 (Term.dest_Type (ProofContext.read_tyname ctxt c))
|
25142
|
237 |
|> syn ? Sign.base_name
|
24695
|
238 |
|> ML_Syntax.print_string));
|
24574
|
239 |
|
25142
|
240 |
val _ = inline_antiq "type_name" (type_ false);
|
|
241 |
val _ = inline_antiq "type_syntax" (type_ true);
|
|
242 |
|
24574
|
243 |
fun const syn = (Scan.state >> Context.proof_of) -- Scan.lift Args.name >> (fn (ctxt, c) =>
|
25334
|
244 |
#1 (Term.dest_Const (ProofContext.read_const' ctxt c))
|
24574
|
245 |
|> syn ? ProofContext.const_syntax_name ctxt
|
|
246 |
|> ML_Syntax.print_string);
|
|
247 |
|
|
248 |
val _ = inline_antiq "const_name" (const false);
|
|
249 |
val _ = inline_antiq "const_syntax" (const true);
|
|
250 |
|
25142
|
251 |
in val _ = () end;
|
24574
|
252 |
|
|
253 |
|
|
254 |
(** fact references **)
|
|
255 |
|
|
256 |
fun thm name = ProofContext.get_thm (the_local_context ()) (Name name);
|
|
257 |
fun thms name = ProofContext.get_thms (the_local_context ()) (Name name);
|
|
258 |
|
|
259 |
|
|
260 |
local
|
|
261 |
|
|
262 |
fun print_interval (FromTo arg) =
|
|
263 |
"PureThy.FromTo " ^ ML_Syntax.print_pair ML_Syntax.print_int ML_Syntax.print_int arg
|
|
264 |
| print_interval (From i) = "PureThy.From " ^ ML_Syntax.print_int i
|
|
265 |
| print_interval (Single i) = "PureThy.Single " ^ ML_Syntax.print_int i;
|
|
266 |
|
|
267 |
fun thm_sel name =
|
|
268 |
Args.thm_sel >> (fn is => "PureThy.NameSelection " ^
|
|
269 |
ML_Syntax.print_pair ML_Syntax.print_string (ML_Syntax.print_list print_interval) (name, is))
|
|
270 |
|| Scan.succeed ("PureThy.Name " ^ ML_Syntax.print_string name);
|
|
271 |
|
|
272 |
fun thm_antiq kind = value_antiq kind
|
|
273 |
(Scan.lift (Args.name :-- thm_sel) >> apsnd (fn sel =>
|
|
274 |
"ProofContext.get_" ^ kind ^ " (ML_Context.the_local_context ()) " ^ ML_Syntax.atomic sel));
|
|
275 |
|
|
276 |
in
|
|
277 |
|
|
278 |
val _ = add_keywords ["(", ")", "-", ","];
|
|
279 |
val _ = thm_antiq "thm";
|
|
280 |
val _ = thm_antiq "thms";
|
|
281 |
|
|
282 |
end;
|
|
283 |
|
24685
|
284 |
val _ = proj_value_antiq "cpat" (Scan.lift Args.name >>
|
|
285 |
(fn name => ("cpat",
|
|
286 |
"Thm.cterm_of (ML_Context.the_context ()) (Syntax.read_term \
|
|
287 |
\(ProofContext.set_mode ProofContext.mode_pattern (ML_Context.the_local_context ()))"
|
|
288 |
^ ML_Syntax.print_string name ^ ")", "")));
|
24574
|
289 |
|
|
290 |
(*final declarations of this structure!*)
|
|
291 |
nonfix >>;
|
|
292 |
fun >> f = change_theory f;
|
|
293 |
|
|
294 |
end;
|
|
295 |
|
|
296 |
structure Basic_ML_Context: BASIC_ML_CONTEXT = ML_Context;
|
|
297 |
open Basic_ML_Context;
|