author | wenzelm |
Tue, 25 Mar 2014 16:54:38 +0100 | |
changeset 56279 | b4d874f6c6be |
parent 56278 | 2576d3a40ed6 |
child 56294 | 85911b8a6868 |
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 ML_CONTEXT = |
|
8 |
sig |
|
9 |
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
|
10 |
val the_global_context: unit -> theory |
24574 | 11 |
val the_local_context: unit -> Proof.context |
39164
e7e12555e763
ML_Context.thm and ML_Context.thms no longer pervasive;
wenzelm
parents:
39140
diff
changeset
|
12 |
val thm: xstring -> thm |
e7e12555e763
ML_Context.thm and ML_Context.thms no longer pervasive;
wenzelm
parents:
39140
diff
changeset
|
13 |
val thms: xstring -> thm list |
26455 | 14 |
val exec: (unit -> unit) -> Context.generic -> Context.generic |
56070 | 15 |
val check_antiquotation: Proof.context -> xstring * Position.T -> string |
56205 | 16 |
type decl = Proof.context -> string * string |
17 |
val add_antiquotation: binding -> (Args.src -> Proof.context -> decl * Proof.context) -> |
|
18 |
theory -> theory |
|
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
19 |
val print_antiquotations: Proof.context -> unit |
56279
b4d874f6c6be
clarified options ML_source_trace and ML_exception_trace (NB: the latter needs to be a system option, since the context is sometimes not available, e.g. for 'theory' command);
wenzelm
parents:
56278
diff
changeset
|
20 |
val source_trace_raw: Config.raw |
b4d874f6c6be
clarified options ML_source_trace and ML_exception_trace (NB: the latter needs to be a system option, since the context is sometimes not available, e.g. for 'theory' command);
wenzelm
parents:
56278
diff
changeset
|
21 |
val source_trace: bool Config.T |
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
22 |
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
|
23 |
Context.generic option -> (ML_Lex.token list * ML_Lex.token list) * Context.generic option |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
24 |
val eval: ML_Compiler.flags -> Position.T -> ML_Lex.token Antiquote.antiquote list -> unit |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
25 |
val eval_file: ML_Compiler.flags -> Path.T -> unit |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
26 |
val eval_source: ML_Compiler.flags -> Symbol_Pos.source -> unit |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
27 |
val eval_in: Proof.context option -> ML_Compiler.flags -> Position.T -> |
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
|
28 |
ML_Lex.token Antiquote.antiquote list -> unit |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
29 |
val eval_source_in: Proof.context option -> ML_Compiler.flags -> Symbol_Pos.source -> unit |
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
|
30 |
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
|
31 |
Context.generic -> Context.generic |
25204 | 32 |
end |
24574 | 33 |
|
34 |
structure ML_Context: ML_CONTEXT = |
|
35 |
struct |
|
36 |
||
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
37 |
(** implicit ML context **) |
24574 | 38 |
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
39 |
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
|
40 |
val the_global_context = Context.theory_of o the_generic_context; |
24574 | 41 |
val the_local_context = Context.proof_of o the_generic_context; |
42 |
||
42360 | 43 |
fun thm name = Proof_Context.get_thm (the_local_context ()) name; |
44 |
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
|
45 |
|
26455 | 46 |
fun exec (e: unit -> unit) context = |
47 |
(case Context.setmp_thread_data (SOME context) (fn () => (e (); Context.thread_data ())) () of |
|
48 |
SOME context' => context' |
|
49 |
| NONE => error "Missing context after execution"); |
|
26416
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
50 |
|
a66f86ef7bb9
added store_thms etc. (formerly in Thy/thm_database.ML);
wenzelm
parents:
26405
diff
changeset
|
51 |
|
24574 | 52 |
|
53 |
(** ML antiquotations **) |
|
54 |
||
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
55 |
(* theory data *) |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
56 |
|
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
57 |
type decl = Proof.context -> string * string; (*final context -> ML env, ML body*) |
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
58 |
structure Antiquotations = Theory_Data |
43560
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
59 |
( |
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
60 |
type T = (Args.src -> Proof.context -> decl * Proof.context) Name_Space.table; |
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
48992
diff
changeset
|
61 |
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
|
62 |
val extend = I; |
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42360
diff
changeset
|
63 |
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
|
64 |
); |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
65 |
|
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
66 |
val get_antiquotations = Antiquotations.get o Proof_Context.theory_of; |
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
67 |
|
56070 | 68 |
fun check_antiquotation ctxt = |
69 |
#1 o Name_Space.check (Context.Proof ctxt) (get_antiquotations ctxt); |
|
70 |
||
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
71 |
fun add_antiquotation name f thy = thy |
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
72 |
|> Antiquotations.map (Name_Space.define (Context.Theory thy) true (name, f) #> snd); |
55743 | 73 |
|
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
74 |
fun print_antiquotations ctxt = |
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
75 |
Pretty.big_list "ML antiquotations:" |
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
76 |
(map (Pretty.mark_str o #1) (Name_Space.markup_table ctxt (get_antiquotations ctxt))) |
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
77 |
|> Pretty.writeln; |
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
78 |
|
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
79 |
fun apply_antiquotation src ctxt = |
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
80 |
let val (src', f) = Args.check_src ctxt (get_antiquotations ctxt) src |
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
81 |
in f src' ctxt end; |
43563 | 82 |
|
24574 | 83 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
84 |
(* parsing and evaluation *) |
24574 | 85 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
86 |
local |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
87 |
|
24574 | 88 |
val antiq = |
56201 | 89 |
Parse.!!! (Parse.position Parse.xname -- Parse.args --| Scan.ahead Parse.eof) |
56029
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
wenzelm
parents:
55828
diff
changeset
|
90 |
>> uncurry Args.src; |
24574 | 91 |
|
48776
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
92 |
val begin_env0 = ML_Lex.tokenize "structure Isabelle =\nstruct\n"; |
52663
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
93 |
|
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
94 |
fun begin_env visible = |
48776
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
95 |
ML_Lex.tokenize |
52663
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
96 |
("structure Isabelle =\nstruct\n\ |
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
97 |
\val ML_context = Context_Position.set_visible " ^ Bool.toString visible ^ |
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
98 |
" (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
|
99 |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
100 |
val end_env = ML_Lex.tokenize "end;"; |
31334 | 101 |
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
|
102 |
|
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
103 |
in |
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
104 |
|
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
105 |
fun eval_antiquotes (ants, pos) opt_context = |
24574 | 106 |
let |
52663
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
107 |
val visible = |
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
108 |
(case opt_context of |
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
109 |
SOME (Context.Proof ctxt) => Context_Position.is_visible ctxt |
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
110 |
| _ => true); |
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
111 |
val opt_ctxt = Option.map (Context.Proof o Context.proof_of) opt_context; |
52663
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
112 |
|
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
113 |
val ((ml_env, ml_body), opt_ctxt') = |
30637
3e3c2cd88cf1
export eval_antiquotes: refined version that operates on ML tokens;
wenzelm
parents:
30614
diff
changeset
|
114 |
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
|
115 |
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
|
116 |
else |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
117 |
let |
53167 | 118 |
val lex = #1 (Keyword.get_lexicons ()); |
119 |
fun no_decl _ = ([], []); |
|
120 |
||
121 |
fun expand (Antiquote.Text tok) ctxt = (K ([], [tok]), ctxt) |
|
55526 | 122 |
| expand (Antiquote.Antiq (ss, {range, ...})) ctxt = |
53167 | 123 |
let |
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
124 |
val (decl, ctxt') = |
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
125 |
apply_antiquotation (Token.read_antiq lex antiq (ss, #1 range)) ctxt; |
53167 | 126 |
val decl' = decl #> pairself (ML_Lex.tokenize #> map (ML_Lex.set_range range)); |
127 |
in (decl', ctxt') end; |
|
128 |
||
27343
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
129 |
val ctxt = |
4b28b80dd1f8
add_antiq: more general notion of ML antiquotation;
wenzelm
parents:
26881
diff
changeset
|
130 |
(case opt_ctxt of |
48992 | 131 |
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
|
132 |
| SOME ctxt => Context.proof_of ctxt); |
27378 | 133 |
|
53167 | 134 |
val (decls, ctxt') = fold_map expand ants ctxt; |
52663
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
135 |
val (ml_env, ml_body) = |
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
136 |
decls |> map (fn decl => decl ctxt') |> split_list |> pairself flat; |
6e71d43775e5
retain compile-time context visibility, which is particularly important for "apply tactic";
wenzelm
parents:
50201
diff
changeset
|
137 |
in ((begin_env visible @ ml_env, ml_body), SOME (Context.Proof ctxt')) end; |
48776
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents:
47005
diff
changeset
|
138 |
in ((ml_env @ end_env, ml_body), opt_ctxt') end; |
24574 | 139 |
|
56279
b4d874f6c6be
clarified options ML_source_trace and ML_exception_trace (NB: the latter needs to be a system option, since the context is sometimes not available, e.g. for 'theory' command);
wenzelm
parents:
56278
diff
changeset
|
140 |
val source_trace_raw = Config.declare "ML_source_trace" (fn _ => Config.Bool false); |
b4d874f6c6be
clarified options ML_source_trace and ML_exception_trace (NB: the latter needs to be a system option, since the context is sometimes not available, e.g. for 'theory' command);
wenzelm
parents:
56278
diff
changeset
|
141 |
val source_trace = Config.bool source_trace_raw; |
24574 | 142 |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
143 |
fun eval flags pos ants = |
24574 | 144 |
let |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
145 |
val non_verbose = {SML = #SML flags, verbose = false}; |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
146 |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
147 |
(*prepare source text*) |
41485
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
148 |
val ((env, body), env_ctxt) = eval_antiquotes (ants, pos) (Context.thread_data ()); |
31334 | 149 |
val _ = |
41485
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
150 |
(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
|
151 |
SOME ctxt => |
56279
b4d874f6c6be
clarified options ML_source_trace and ML_exception_trace (NB: the latter needs to be a system option, since the context is sometimes not available, e.g. for 'theory' command);
wenzelm
parents:
56278
diff
changeset
|
152 |
if Config.get ctxt source_trace then |
41485
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
153 |
Context_Position.if_visible ctxt |
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
154 |
tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body]) |
41375 | 155 |
else () |
156 |
| NONE => ()); |
|
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
157 |
|
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
158 |
(*prepare static ML environment*) |
41485
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
159 |
val _ = |
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
160 |
Context.setmp_thread_data |
6c0de045d127
ML_trace: observe context visibility flag (import for Latex mode, for example);
wenzelm
parents:
41375
diff
changeset
|
161 |
(Option.map (Context.mapping I (Context_Position.set_visible false)) env_ctxt) |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
162 |
(fn () => (ML_Compiler.eval non_verbose Position.none env; Context.thread_data ())) () |
31325
700951b53d21
moved local ML environment to separate module ML_Env;
wenzelm
parents:
30683
diff
changeset
|
163 |
|> (fn NONE => () | SOME context' => Context.>> (ML_Env.inherit context')); |
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
164 |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
165 |
val _ = ML_Compiler.eval flags pos body; |
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
166 |
val _ = ML_Compiler.eval non_verbose Position.none reset_env; |
28270
7ada24ebea94
explicit handling of ML environment within generic context;
wenzelm
parents:
27895
diff
changeset
|
167 |
in () end; |
24574 | 168 |
|
169 |
end; |
|
170 |
||
171 |
||
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30643
diff
changeset
|
172 |
(* derived versions *) |
24574 | 173 |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
174 |
fun eval_file flags path = |
56203
76c72f4d0667
clarified bootstrap process: switch to ML with context and antiquotations earlier;
wenzelm
parents:
56201
diff
changeset
|
175 |
let val pos = Path.position path |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
176 |
in eval flags pos (ML_Lex.read pos (File.read path)) end; |
56203
76c72f4d0667
clarified bootstrap process: switch to ML with context and antiquotations earlier;
wenzelm
parents:
56201
diff
changeset
|
177 |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
178 |
fun eval_source flags source = |
56278
2576d3a40ed6
separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents:
56275
diff
changeset
|
179 |
eval flags (#pos source) (ML_Lex.read_source (#SML flags) source); |
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
|
180 |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
181 |
fun eval_in ctxt flags pos ants = |
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
182 |
Context.setmp_thread_data (Option.map Context.Proof ctxt) |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
183 |
(fn () => eval flags pos ants) (); |
24574 | 184 |
|
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
185 |
fun eval_source_in ctxt flags source = |
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
56032
diff
changeset
|
186 |
Context.setmp_thread_data (Option.map Context.Proof ctxt) |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
187 |
(fn () => eval_source flags source) (); |
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
|
188 |
|
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
|
189 |
fun expression pos bind body ants = |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
190 |
exec (fn () => eval {SML = false, verbose = false} pos |
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
|
191 |
(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
|
192 |
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
|
193 |
|
24574 | 194 |
end; |
195 |
||
56229
f61eaab6bec3
tuned error, according to "use" in General/secure.ML;
wenzelm
parents:
56205
diff
changeset
|
196 |
fun use s = |
56275
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
wenzelm
parents:
56229
diff
changeset
|
197 |
ML_Context.eval_file {SML = false, verbose = true} (Path.explode s) |
56229
f61eaab6bec3
tuned error, according to "use" in General/secure.ML;
wenzelm
parents:
56205
diff
changeset
|
198 |
handle ERROR msg => (writeln msg; error "ML error"); |
f61eaab6bec3
tuned error, according to "use" in General/secure.ML;
wenzelm
parents:
56205
diff
changeset
|
199 |