| author | wenzelm |
| Sun, 31 May 2009 17:45:53 +0200 | |
| changeset 31320 | 72eeb1b4e006 |
| parent 30683 | e8ac1f9d9469 |
| child 31325 | 700951b53d21 |
| 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 |
| 24574 | 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 the_generic_context: unit -> Context.generic |
|
|
26432
095e448b95a0
renamed ML_Context.the_context to ML_Context.the_global_context;
wenzelm
parents:
26416
diff
changeset
|
19 |
val the_global_context: unit -> theory |
| 24574 | 20 |
val the_local_context: unit -> Proof.context |
| 26455 | 21 |
val exec: (unit -> unit) -> Context.generic -> Context.generic |
| 30574 | 22 |
val inherit_env: Context.generic -> Context.generic -> Context.generic |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
23 |
val name_space: ML_Name_Space.T |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
24 |
val local_context: use_context |
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
25 |
val stored_thms: thm list ref |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
26 |
val ml_store_thm: string * thm -> unit |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
27 |
val ml_store_thms: string * thm list -> unit |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
28 |
type antiq = |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
29 |
{struct_name: string, background: Proof.context} ->
|
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
30 |
(Proof.context -> string * string) * Proof.context |
| 30513 | 31 |
val add_antiq: string -> (Position.T -> antiq context_parser) -> unit |
| 24574 | 32 |
val trace: bool ref |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
33 |
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
|
34 |
Context.generic option -> (ML_Lex.token list * ML_Lex.token list) * Context.generic option |
| 30574 | 35 |
val eval: bool -> Position.T -> Symbol_Pos.text -> unit |
| 26455 | 36 |
val eval_file: Path.T -> unit |
| 30574 | 37 |
val eval_in: Proof.context option -> bool -> Position.T -> Symbol_Pos.text -> unit |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
38 |
val evaluate: Proof.context -> bool -> string * (unit -> 'a) option ref -> string -> 'a |
| 26455 | 39 |
val expression: Position.T -> string -> string -> string -> Context.generic -> Context.generic |
| 25204 | 40 |
end |
| 24574 | 41 |
|
42 |
structure ML_Context: ML_CONTEXT = |
|
43 |
struct |
|
44 |
||
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
45 |
(** implicit ML context **) |
| 24574 | 46 |
|
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
47 |
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
|
48 |
val the_global_context = Context.theory_of o the_generic_context; |
| 24574 | 49 |
val the_local_context = Context.proof_of o the_generic_context; |
50 |
||
| 26455 | 51 |
fun exec (e: unit -> unit) context = |
52 |
(case Context.setmp_thread_data (SOME context) (fn () => (e (); Context.thread_data ())) () of |
|
53 |
SOME context' => context' |
|
54 |
| NONE => error "Missing context after execution"); |
|
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
55 |
|
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
56 |
|
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
57 |
(* ML name space *) |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
58 |
|
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
59 |
structure ML_Env = GenericDataFun |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
60 |
( |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
61 |
type T = |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
62 |
ML_Name_Space.valueVal Symtab.table * |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
63 |
ML_Name_Space.typeVal Symtab.table * |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
64 |
ML_Name_Space.fixityVal Symtab.table * |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
65 |
ML_Name_Space.structureVal Symtab.table * |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
66 |
ML_Name_Space.signatureVal Symtab.table * |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
67 |
ML_Name_Space.functorVal Symtab.table; |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
68 |
val empty = (Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty); |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
69 |
val extend = I; |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
70 |
fun merge _ |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
71 |
((val1, type1, fixity1, structure1, signature1, functor1), |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
72 |
(val2, type2, fixity2, structure2, signature2, functor2)) : T = |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
73 |
(Symtab.merge (K true) (val1, val2), |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
74 |
Symtab.merge (K true) (type1, type2), |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
75 |
Symtab.merge (K true) (fixity1, fixity2), |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
76 |
Symtab.merge (K true) (structure1, structure2), |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
77 |
Symtab.merge (K true) (signature1, signature2), |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
78 |
Symtab.merge (K true) (functor1, functor2)); |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
79 |
); |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
80 |
|
| 30574 | 81 |
val inherit_env = ML_Env.put o ML_Env.get; |
| 28279 | 82 |
|
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
83 |
val name_space: ML_Name_Space.T = |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
84 |
let |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
85 |
fun lookup sel1 sel2 name = |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
86 |
Context.thread_data () |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
87 |
|> (fn NONE => NONE | SOME context => Symtab.lookup (sel1 (ML_Env.get context)) name) |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
88 |
|> (fn NONE => sel2 ML_Name_Space.global name | some => some); |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
89 |
|
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
90 |
fun all sel1 sel2 () = |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
91 |
Context.thread_data () |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
92 |
|> (fn NONE => [] | SOME context => Symtab.dest (sel1 (ML_Env.get context))) |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
93 |
|> append (sel2 ML_Name_Space.global ()) |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
94 |
|> sort_distinct (string_ord o pairself #1); |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
95 |
|
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
96 |
fun enter ap1 sel2 entry = |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
97 |
if is_some (Context.thread_data ()) then |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
98 |
Context.>> (ML_Env.map (ap1 (Symtab.update entry))) |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
99 |
else sel2 ML_Name_Space.global entry; |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
100 |
in |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
101 |
{lookupVal = lookup #1 #lookupVal,
|
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
102 |
lookupType = lookup #2 #lookupType, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
103 |
lookupFix = lookup #3 #lookupFix, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
104 |
lookupStruct = lookup #4 #lookupStruct, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
105 |
lookupSig = lookup #5 #lookupSig, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
106 |
lookupFunct = lookup #6 #lookupFunct, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
107 |
enterVal = enter (fn h => fn (a, b, c, d, e, f) => (h a, b, c, d, e, f)) #enterVal, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
108 |
enterType = enter (fn h => fn (a, b, c, d, e, f) => (a, h b, c, d, e, f)) #enterType, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
109 |
enterFix = enter (fn h => fn (a, b, c, d, e, f) => (a, b, h c, d, e, f)) #enterFix, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
110 |
enterStruct = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, h d, e, f)) #enterStruct, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
111 |
enterSig = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, h e, f)) #enterSig, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
112 |
enterFunct = enter (fn h => fn (a, b, c, d, e, f) => (a, b, c, d, e, h f)) #enterFunct, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
113 |
allVal = all #1 #allVal, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
114 |
allType = all #2 #allType, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
115 |
allFix = all #3 #allFix, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
116 |
allStruct = all #4 #allStruct, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
117 |
allSig = all #5 #allSig, |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
118 |
allFunct = all #6 #allFunct} |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
119 |
end; |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
120 |
|
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
121 |
val local_context: use_context = |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
122 |
{tune_source = ML_Parse.fix_ints,
|
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
123 |
name_space = name_space, |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
124 |
str_of_pos = Position.str_of oo Position.line_file, |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
125 |
print = writeln, |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
126 |
error = error}; |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
127 |
|
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
128 |
|
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
129 |
(* theorem bindings *) |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
130 |
|
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
131 |
val stored_thms: thm list ref = ref []; |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
132 |
|
|
26492
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
133 |
fun ml_store sel (name, ths) = |
|
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
134 |
let |
| 29579 | 135 |
val ths' = Context.>>> (Context.map_theory_result |
136 |
(PureThy.store_thms (Binding.name name, ths))); |
|
|
26492
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
137 |
val _ = |
|
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
138 |
if name = "" then () |
|
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
139 |
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
|
140 |
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
|
141 |
else setmp stored_thms ths' (fn () => |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
142 |
use_text local_context (0, "") true |
|
26492
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
143 |
("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
|
144 |
in () end; |
| 24574 | 145 |
|
|
26492
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
146 |
val ml_store_thms = ml_store ""; |
|
6e87cc839632
removed obsolete store_thm(s), cf. functional versions in pure_thy.ML;
wenzelm
parents:
26473
diff
changeset
|
147 |
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
|
148 |
|
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
149 |
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
|
150 |
fun bind_thms (name, thms) = ml_store_thms (name, map Drule.standard thms); |
| 24574 | 151 |
|
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
152 |
fun thm name = ProofContext.get_thm (the_local_context ()) name; |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
153 |
fun thms name = ProofContext.get_thms (the_local_context ()) name; |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
154 |
|
| 24574 | 155 |
|
156 |
||
157 |
(** ML antiquotations **) |
|
158 |
||
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
159 |
(* antiquotation commands *) |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
160 |
|
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
161 |
type antiq = |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
162 |
{struct_name: string, background: Proof.context} ->
|
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
163 |
(Proof.context -> string * string) * Proof.context; |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
164 |
|
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
165 |
local |
| 24574 | 166 |
|
| 30513 | 167 |
val global_parsers = ref (Symtab.empty: (Position.T -> antiq context_parser) Symtab.table); |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
168 |
|
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
169 |
in |
| 24574 | 170 |
|
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
171 |
fun add_antiq name scan = CRITICAL (fn () => |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
172 |
change global_parsers (fn tab => |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
173 |
(if not (Symtab.defined tab name) then () |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
174 |
else warning ("Redefined ML antiquotation: " ^ quote name);
|
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
175 |
Symtab.update (name, scan) tab))); |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
176 |
|
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
177 |
fun antiquotation src ctxt = |
| 24574 | 178 |
let val ((name, _), pos) = Args.dest_src src in |
179 |
(case Symtab.lookup (! global_parsers) name of |
|
180 |
NONE => error ("Unknown ML antiquotation command: " ^ quote name ^ Position.str_of pos)
|
|
|
28412
0608c04858c7
back to plain Position.report for regular references;
wenzelm
parents:
28407
diff
changeset
|
181 |
| SOME scan => (Position.report (Markup.ML_antiq name) pos; |
| 27895 | 182 |
Args.context_syntax "ML antiquotation" (scan pos) src ctxt)) |
| 24574 | 183 |
end; |
184 |
||
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
185 |
end; |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
186 |
|
| 24574 | 187 |
|
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
188 |
(* parsing and evaluation *) |
| 24574 | 189 |
|
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
190 |
local |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
191 |
|
| 24574 | 192 |
structure P = OuterParse; |
| 30588 | 193 |
structure T = OuterLex; |
| 24574 | 194 |
|
195 |
val antiq = |
|
|
27809
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27781
diff
changeset
|
196 |
P.!!! (P.position P.xname -- Args.parse --| Scan.ahead P.eof) |
| 24574 | 197 |
>> (fn ((x, pos), y) => Args.src ((x, y), pos)); |
198 |
||
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
199 |
val struct_name = "Isabelle"; |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
200 |
val begin_env = ML_Lex.tokenize ("structure " ^ struct_name ^ " =\nstruct\n");
|
|
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
201 |
val end_env = ML_Lex.tokenize "end;"; |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
202 |
|
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
203 |
in |
|
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
204 |
|
|
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
205 |
fun eval_antiquotes (ants, pos) opt_context = |
| 24574 | 206 |
let |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
207 |
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
|
208 |
val ((ml_env, ml_body), opt_ctxt') = |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
209 |
if forall Antiquote.is_text ants |
|
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
210 |
then (([], map (fn Antiquote.Text tok => tok) ants), opt_ctxt) |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
211 |
else |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
212 |
let |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
213 |
val ctxt = |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
214 |
(case opt_ctxt of |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
215 |
NONE => error ("No context -- cannot expand ML antiquotations" ^ Position.str_of pos)
|
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
216 |
| SOME ctxt => Context.proof_of ctxt); |
| 27378 | 217 |
|
| 27359 | 218 |
val lex = #1 (OuterKeyword.get_lexicons ()); |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
219 |
fun no_decl _ = ([], []); |
| 27378 | 220 |
|
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
221 |
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
|
222 |
| expand (Antiquote.Antiq (ss, range)) (scope, background) = |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
223 |
let |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
224 |
val context = Stack.top scope; |
|
30683
e8ac1f9d9469
datatype antiquote: maintain original Position.range, which is eventually attached to the resulting ML tokens;
wenzelm
parents:
30672
diff
changeset
|
225 |
val (f, context') = antiquotation (T.read_antiq lex antiq (ss, #1 range)) context; |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
226 |
val (decl, background') = f {background = background, struct_name = struct_name};
|
|
30683
e8ac1f9d9469
datatype antiquote: maintain original Position.range, which is eventually attached to the resulting ML tokens;
wenzelm
parents:
30672
diff
changeset
|
227 |
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
|
228 |
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
|
229 |
| expand (Antiquote.Open _) (scope, background) = |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
230 |
(no_decl, (Stack.push scope, background)) |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
231 |
| expand (Antiquote.Close _) (scope, background) = |
|
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
232 |
(no_decl, (Stack.pop scope, background)); |
| 24574 | 233 |
|
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
234 |
val (decls, (_, ctxt')) = fold_map expand ants (Stack.init ctxt, ctxt); |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
235 |
val ml = decls |> map (fn decl => decl ctxt') |> split_list |> pairself flat; |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
236 |
in (ml, SOME (Context.Proof ctxt')) end; |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
237 |
in ((begin_env @ ml_env @ end_env, ml_body), opt_ctxt') end; |
| 24574 | 238 |
|
239 |
val trace = ref false; |
|
240 |
||
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
241 |
fun eval verbose pos txt = |
| 24574 | 242 |
let |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
243 |
fun eval_raw p = use_text local_context |
|
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
244 |
(the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p)); |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
245 |
|
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
246 |
(*prepare source text*) |
| 27878 | 247 |
val _ = Position.report Markup.ML_source pos; |
| 30643 | 248 |
val ants = ML_Lex.read_antiq (Symbol_Pos.explode (txt, pos), pos); |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
249 |
val ((env, body), env_ctxt) = eval_antiquotes (ants, pos) (Context.thread_data ()) |
|
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
250 |
|>> pairself (implode o map ML_Lex.text_of); |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
251 |
val _ = if ! trace then tracing (cat_lines [env, body]) else (); |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
252 |
|
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
253 |
(*prepare static ML environment*) |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
254 |
val _ = Context.setmp_thread_data env_ctxt |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
255 |
(fn () => (eval_raw Position.none false env; Context.thread_data ())) () |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
256 |
|> (fn NONE => () | SOME context' => Context.>> (inherit_env context')); |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
257 |
|
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
258 |
(*eval ML*) |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
259 |
val _ = eval_raw pos verbose body; |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
260 |
|
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
261 |
(*reset static ML environment*) |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
262 |
val _ = eval_raw Position.none false "structure Isabelle = struct end"; |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
263 |
in () end; |
| 24574 | 264 |
|
265 |
end; |
|
266 |
||
267 |
||
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
268 |
(* derived versions *) |
| 24574 | 269 |
|
| 26881 | 270 |
fun eval_file path = eval true (Path.position path) (File.read path); |
| 24574 | 271 |
|
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
272 |
fun eval_in ctxt verbose pos txt = |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
273 |
Context.setmp_thread_data (Option.map Context.Proof ctxt) (fn () => eval verbose pos txt) (); |
| 25204 | 274 |
|
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
275 |
fun evaluate ctxt verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () => |
| 25204 | 276 |
let |
277 |
val _ = r := NONE; |
|
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
278 |
val _ = Context.setmp_thread_data (SOME (Context.Proof ctxt)) (fn () => |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
279 |
eval verbose Position.none ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))")) ();
|
| 25204 | 280 |
in (case ! r of NONE => error ("Bad evaluation for " ^ ref_name) | SOME e => e) end) ();
|
| 24574 | 281 |
|
| 26455 | 282 |
fun expression pos bind body txt = |
283 |
exec (fn () => eval false pos |
|
284 |
("Context.set_thread_data (SOME (let " ^ bind ^ " = " ^ txt ^ " in " ^ body ^
|
|
285 |
" end (ML_Context.the_generic_context ())));")); |
|
286 |
||
| 24574 | 287 |
end; |
288 |
||
289 |
structure Basic_ML_Context: BASIC_ML_CONTEXT = ML_Context; |
|
290 |
open Basic_ML_Context; |