author | nipkow |
Fri, 30 Nov 2012 17:12:01 +0100 | |
changeset 50303 | 5c4c35321e87 |
parent 50201 | c26369c9eda6 |
child 52663 | 6e71d43775e5 |
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:
26405
diff
changeset
|
9 |
val bind_thm: string * thm -> unit |
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
10 |
val bind_thms: string * thm list -> unit |
25204 | 11 |
end |
24574 | 12 |
|
13 |
signature ML_CONTEXT = |
|
14 |
sig |
|
15 |
include BASIC_ML_CONTEXT |
|
16 |
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
|
17 |
val the_global_context: unit -> theory |
24574 | 18 |
val the_local_context: unit -> Proof.context |
39164
e7e12555e763
ML_Context.thm and ML_Context.thms no longer pervasive;
wenzelm
parents:
39140
diff
changeset
|
19 |
val thm: xstring -> thm |
e7e12555e763
ML_Context.thm and ML_Context.thms no longer pervasive;
wenzelm
parents:
39140
diff
changeset
|
20 |
val thms: xstring -> thm list |
26455 | 21 |
val exec: (unit -> unit) -> Context.generic -> Context.generic |
39140
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
22 |
val get_stored_thms: unit -> thm list |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
23 |
val get_stored_thm: unit -> thm |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
24 |
val ml_store_thms: string * thm list -> unit |
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
25 |
val ml_store_thm: string * thm -> unit |
35019
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
wenzelm
parents:
33222
diff
changeset
|
26 |
type antiq = Proof.context -> (Proof.context -> string * string) * Proof.context |
43560
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
27 |
val add_antiq: binding -> (Position.T -> antiq context_parser) -> theory -> theory |
43563 | 28 |
val intern_antiq: theory -> xstring -> string |
29 |
val defined_antiq: theory -> string -> bool |
|
41375 | 30 |
val trace_raw: Config.raw |
31 |
val trace: bool Config.T |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
32 |
val eval_antiquotes: ML_Lex.token Antiquote.antiquote list * Position.T -> |
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
33 |
Context.generic option -> (ML_Lex.token list * ML_Lex.token list) * Context.generic option |
37198
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
34 |
val eval: bool -> Position.T -> ML_Lex.token Antiquote.antiquote list -> unit |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
35 |
val eval_text: bool -> Position.T -> Symbol_Pos.text -> unit |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
36 |
val eval_in: Proof.context option -> bool -> Position.T -> |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
37 |
ML_Lex.token Antiquote.antiquote list -> unit |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
38 |
val eval_text_in: Proof.context option -> bool -> Position.T -> Symbol_Pos.text -> unit |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
39 |
val expression: Position.T -> string -> string -> ML_Lex.token Antiquote.antiquote list -> |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
40 |
Context.generic -> Context.generic |
25204 | 41 |
end |
24574 | 42 |
|
43 |
structure ML_Context: ML_CONTEXT = |
|
44 |
struct |
|
45 |
||
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
46 |
(** implicit ML context **) |
24574 | 47 |
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
48 |
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
|
49 |
val the_global_context = Context.theory_of o the_generic_context; |
24574 | 50 |
val the_local_context = Context.proof_of o the_generic_context; |
51 |
||
42360 | 52 |
fun thm name = Proof_Context.get_thm (the_local_context ()) name; |
53 |
fun thms name = Proof_Context.get_thms (the_local_context ()) name; |
|
39164
e7e12555e763
ML_Context.thm and ML_Context.thms no longer pervasive;
wenzelm
parents:
39140
diff
changeset
|
54 |
|
26455 | 55 |
fun exec (e: unit -> unit) context = |
56 |
(case Context.setmp_thread_data (SOME context) (fn () => (e (); Context.thread_data ())) () of |
|
57 |
SOME context' => context' |
|
58 |
| NONE => error "Missing context after execution"); |
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
59 |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
60 |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
61 |
(* theorem bindings *) |
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
62 |
|
39140
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
63 |
structure Stored_Thms = Theory_Data |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
64 |
( |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
65 |
type T = thm list; |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
66 |
val empty = []; |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
67 |
fun extend _ = []; |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
68 |
fun merge _ = []; |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
69 |
); |
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
70 |
|
39140
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
71 |
fun get_stored_thms () = Stored_Thms.get (the_global_context ()); |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
72 |
val get_stored_thm = hd o get_stored_thms; |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
73 |
|
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
74 |
fun ml_store get (name, ths) = |
26492
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
75 |
let |
29579 | 76 |
val ths' = Context.>>> (Context.map_theory_result |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39507
diff
changeset
|
77 |
(Global_Theory.store_thms (Binding.name name, ths))); |
39140
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
78 |
val _ = Context.>> (Context.map_theory (Stored_Thms.put ths')); |
26492
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
79 |
val _ = |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
80 |
if name = "" then () |
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
81 |
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
|
82 |
error ("Cannot bind theorem(s) " ^ quote name ^ " as ML value") |
39140
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
83 |
else |
31334 | 84 |
ML_Compiler.eval true Position.none |
39140
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
85 |
(ML_Lex.tokenize ("val " ^ name ^ " = " ^ get ^ " ();")); |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
86 |
val _ = Context.>> (Context.map_theory (Stored_Thms.put [])); |
26492
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
87 |
in () end; |
24574 | 88 |
|
39140
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
89 |
val ml_store_thms = ml_store "ML_Context.get_stored_thms"; |
8a2fb4ce1f39
ML_Context.ml_store_thm(s): operate on thread context instead of Unsynchronized.ref;
wenzelm
parents:
38931
diff
changeset
|
90 |
fun ml_store_thm (name, th) = ml_store "ML_Context.get_stored_thm" (name, [th]); |
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
91 |
|
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
35019
diff
changeset
|
92 |
fun bind_thm (name, thm) = ml_store_thm (name, Drule.export_without_context thm); |
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
35019
diff
changeset
|
93 |
fun bind_thms (name, thms) = ml_store_thms (name, map Drule.export_without_context thms); |
24574 | 94 |
|
95 |
||
96 |
||
97 |
(** ML antiquotations **) |
|
98 |
||
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
99 |
(* antiquotation commands *) |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
100 |
|
35019
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
wenzelm
parents:
33222
diff
changeset
|
101 |
type antiq = Proof.context -> (Proof.context -> string * string) * Proof.context; |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
102 |
|
43560
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
103 |
structure Antiq_Parsers = Theory_Data |
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
104 |
( |
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
105 |
type T = (Position.T -> antiq context_parser) Name_Space.table; |
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
48992
diff
changeset
|
106 |
val empty : T = Name_Space.empty_table Markup.ML_antiquotationN; |
43560
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
107 |
val extend = I; |
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
108 |
fun merge data : T = Name_Space.merge_tables data; |
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
109 |
); |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
110 |
|
43560
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
111 |
fun add_antiq name scan thy = thy |
47005
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
45666
diff
changeset
|
112 |
|> Antiq_Parsers.map (Name_Space.define (Context.Theory thy) true (name, scan) #> snd); |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
113 |
|
43563 | 114 |
val intern_antiq = Name_Space.intern o #1 o Antiq_Parsers.get; |
115 |
val defined_antiq = Symtab.defined o #2 o Antiq_Parsers.get; |
|
116 |
||
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
117 |
fun antiquotation src ctxt = |
43560
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
118 |
let |
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
119 |
val thy = Proof_Context.theory_of ctxt; |
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
120 |
val ((xname, _), pos) = Args.dest_src src; |
47005
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
45666
diff
changeset
|
121 |
val (_, scan) = Name_Space.check (Context.Proof ctxt) (Antiq_Parsers.get thy) (xname, pos); |
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
48992
diff
changeset
|
122 |
in Args.context_syntax Markup.ML_antiquotationN (scan pos) src ctxt end; |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
123 |
|
24574 | 124 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
125 |
(* parsing and evaluation *) |
24574 | 126 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
127 |
local |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
128 |
|
24574 | 129 |
val antiq = |
36950 | 130 |
Parse.!!! (Parse.position Parse.xname -- Args.parse --| Scan.ahead Parse.eof) |
24574 | 131 |
>> (fn ((x, pos), y) => Args.src ((x, y), pos)); |
132 |
||
48776
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
133 |
val begin_env0 = ML_Lex.tokenize "structure Isabelle =\nstruct\n"; |
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
134 |
val begin_env = |
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
135 |
ML_Lex.tokenize |
48781 | 136 |
"structure Isabelle =\nstruct\nval ML_context = ML_Context.the_local_context ();\n"; |
48776
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
137 |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
138 |
val end_env = ML_Lex.tokenize "end;"; |
31334 | 139 |
val reset_env = ML_Lex.tokenize "structure Isabelle = struct end"; |
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
140 |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
141 |
in |
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
142 |
|
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
143 |
fun eval_antiquotes (ants, pos) opt_context = |
24574 | 144 |
let |
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
145 |
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:
26881
diff
changeset
|
146 |
val ((ml_env, ml_body), opt_ctxt') = |
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
147 |
if forall Antiquote.is_text ants |
48776
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
148 |
then ((begin_env0, map (fn Antiquote.Text tok => tok) ants), opt_ctxt) |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
149 |
else |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
150 |
let |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
151 |
val ctxt = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
152 |
(case opt_ctxt of |
48992 | 153 |
NONE => error ("No context -- cannot expand ML antiquotations" ^ Position.here pos) |
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
154 |
| SOME ctxt => Context.proof_of ctxt); |
27378 | 155 |
|
36950 | 156 |
val lex = #1 (Keyword.get_lexicons ()); |
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
157 |
fun no_decl _ = ([], []); |
27378 | 158 |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
159 |
fun expand (Antiquote.Text tok) state = (K ([], [tok]), state) |
30683
e8ac1f9d9469
datatype antiquote: maintain original Position.range, which is eventually attached to the resulting ML tokens;
wenzelm
parents:
30672
diff
changeset
|
160 |
| expand (Antiquote.Antiq (ss, range)) (scope, background) = |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
161 |
let |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
162 |
val context = Stack.top scope; |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
163 |
val (f, context') = |
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
164 |
antiquotation (Token.read_antiq lex antiq (ss, #1 range)) context; |
35019
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
wenzelm
parents:
33222
diff
changeset
|
165 |
val (decl, background') = f background; |
30683
e8ac1f9d9469
datatype antiquote: maintain original Position.range, which is eventually attached to the resulting ML tokens;
wenzelm
parents:
30672
diff
changeset
|
166 |
val decl' = decl #> pairself (ML_Lex.tokenize #> map (ML_Lex.set_range range)); |
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
167 |
in (decl', (Stack.map_top (K context') scope, background')) end |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
168 |
| expand (Antiquote.Open _) (scope, background) = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
169 |
(no_decl, (Stack.push scope, background)) |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
170 |
| expand (Antiquote.Close _) (scope, background) = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
171 |
(no_decl, (Stack.pop scope, background)); |
24574 | 172 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
173 |
val (decls, (_, ctxt')) = fold_map expand ants (Stack.init ctxt, ctxt); |
48776
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
174 |
val (ml_env, ml_body) = decls |> map (fn decl => decl ctxt') |> split_list |> pairself flat; |
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
175 |
in ((begin_env @ ml_env, ml_body), SOME (Context.Proof ctxt')) end; |
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
176 |
in ((ml_env @ end_env, ml_body), opt_ctxt') end; |
24574 | 177 |
|
41375 | 178 |
val trace_raw = Config.declare "ML_trace" (fn _ => Config.Bool false); |
179 |
val trace = Config.bool trace_raw; |
|
24574 | 180 |
|
37198
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
181 |
fun eval verbose pos ants = |
24574 | 182 |
let |
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
183 |
(*prepare source text*) |
41485
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
184 |
val ((env, body), env_ctxt) = eval_antiquotes (ants, pos) (Context.thread_data ()); |
31334 | 185 |
val _ = |
41485
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
186 |
(case Option.map Context.proof_of env_ctxt of |
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
187 |
SOME ctxt => |
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
188 |
if Config.get ctxt trace then |
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
189 |
Context_Position.if_visible ctxt |
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
190 |
tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body]) |
41375 | 191 |
else () |
192 |
| NONE => ()); |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
193 |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
194 |
(*prepare static ML environment*) |
41485
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
195 |
val _ = |
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
196 |
Context.setmp_thread_data |
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
197 |
(Option.map (Context.mapping I (Context_Position.set_visible false)) env_ctxt) |
31334 | 198 |
(fn () => (ML_Compiler.eval false Position.none env; Context.thread_data ())) () |
31325
700951b53d21
moved local ML environment to separate module ML_Env;
wenzelm
parents:
30683
diff
changeset
|
199 |
|> (fn NONE => () | SOME context' => Context.>> (ML_Env.inherit context')); |
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
200 |
|
31334 | 201 |
val _ = ML_Compiler.eval verbose pos body; |
202 |
val _ = ML_Compiler.eval false Position.none reset_env; |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
203 |
in () end; |
24574 | 204 |
|
205 |
end; |
|
206 |
||
207 |
||
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
208 |
(* derived versions *) |
24574 | 209 |
|
37198
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
210 |
fun eval_text verbose pos txt = eval verbose pos (ML_Lex.read pos txt); |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
211 |
|
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
212 |
fun eval_in ctxt verbose pos ants = |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
213 |
Context.setmp_thread_data (Option.map Context.Proof ctxt) (fn () => eval verbose pos ants) (); |
24574 | 214 |
|
37198
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
215 |
fun eval_text_in ctxt verbose pos txt = |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
216 |
Context.setmp_thread_data (Option.map Context.Proof ctxt) (fn () => eval_text verbose pos txt) (); |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
217 |
|
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
218 |
fun expression pos bind body ants = |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
219 |
exec (fn () => eval false pos |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
220 |
(ML_Lex.read Position.none ("Context.set_thread_data (SOME (let " ^ bind ^ " = ") @ ants @ |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
221 |
ML_Lex.read Position.none (" in " ^ body ^ " end (ML_Context.the_generic_context ())));"))); |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
222 |
|
24574 | 223 |
end; |
224 |
||
225 |
structure Basic_ML_Context: BASIC_ML_CONTEXT = ML_Context; |
|
226 |
open Basic_ML_Context; |