| author | wenzelm |
| Tue, 24 Mar 2009 11:57:41 +0100 | |
| changeset 30683 | e8ac1f9d9469 |
| parent 30681 | 27ee3f4ea99c |
| child 30704 | d6d4828e74a2 |
| permissions | -rw-r--r-- |
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Pure/ML/ml_test.ML |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
2 |
Author: Makarius |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
3 |
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
4 |
Test of advanced ML compiler invocation in Poly/ML 5.3. |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
5 |
*) |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
6 |
|
| 30646 | 7 |
signature ML_TEST = |
8 |
sig |
|
| 30668 | 9 |
val get_result: Proof.context -> PolyML.parseTree list |
| 30680 | 10 |
val eval: bool -> Position.T -> ML_Lex.token list -> unit |
| 30646 | 11 |
end |
12 |
||
13 |
structure ML_Test: ML_TEST = |
|
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
14 |
struct |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
15 |
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
16 |
(* eval ML source tokens *) |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
17 |
|
| 30646 | 18 |
structure Result = GenericDataFun |
19 |
( |
|
| 30668 | 20 |
type T = PolyML.parseTree list; |
21 |
val empty = []; |
|
22 |
fun extend _ = []; |
|
23 |
fun merge _ _ = []; |
|
| 30646 | 24 |
); |
25 |
||
26 |
val get_result = Result.get o Context.Proof; |
|
27 |
||
| 30680 | 28 |
|
29 |
fun use_text ({name_space = space, print, error, ...}: use_context) verbose pos toks =
|
|
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
30 |
let |
| 30680 | 31 |
(* input *) |
32 |
||
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
33 |
val input = toks |> map (fn tok => |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
34 |
(serial (), (String.explode (ML_Lex.text_of tok), ML_Lex.pos_of tok))); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
35 |
val index_pos = Inttab.lookup (Inttab.make (map (apsnd snd) input)); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
36 |
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
37 |
fun pos_of ({file, startLine = line, startPosition = i, endPosition = j, ...}: location) =
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
38 |
(case (index_pos i, index_pos j) of |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
39 |
(SOME p, SOME q) => Position.encode_range (p, q) |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
40 |
| (SOME p, NONE) => p |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
41 |
| _ => Position.line_file line file); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
42 |
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
43 |
val in_buffer = ref (map (apsnd fst) input); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
44 |
val current_line = ref (the_default 1 (Position.line_of pos)); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
45 |
fun get () = |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
46 |
(case ! in_buffer of |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
47 |
[] => NONE |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
48 |
| (_, []) :: rest => (in_buffer := rest; get ()) |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
49 |
| (i, c :: cs) :: rest => |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
50 |
(in_buffer := (i, cs) :: rest; |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
51 |
if c = #"\n" then current_line := ! current_line + 1 else (); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
52 |
SOME c)); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
53 |
fun get_index () = (case ! in_buffer of [] => 0 | (i, _) :: _ => i); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
54 |
|
| 30680 | 55 |
|
56 |
(* output *) |
|
57 |
||
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
58 |
val out_buffer = ref ([]: string list); |
| 30680 | 59 |
fun output () = implode (rev (! out_buffer)); |
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
60 |
fun put s = out_buffer := s :: ! out_buffer; |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
61 |
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
62 |
fun put_message {message, hard, location, context = _} =
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
63 |
(put (if hard then "Error: " else "Warning: "); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
64 |
put (Pretty.string_of (Pretty.from_ML (pretty_ml message))); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
65 |
put (Position.str_of (pos_of location) ^ "\n")); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
66 |
|
| 30680 | 67 |
|
68 |
(* results *) |
|
69 |
||
70 |
fun apply_result {fixes, types, signatures, structures, functors, values} =
|
|
71 |
let |
|
72 |
fun add_prefix prefix (PrettyBlock (ind, consistent, context, prts)) = |
|
73 |
let |
|
74 |
fun make_prefix context = |
|
75 |
(case get_first (fn ContextParentStructure p => SOME p | _ => NONE) context of |
|
76 |
SOME (name, sub_context) => make_prefix sub_context ^ name ^ "." |
|
77 |
| NONE => prefix); |
|
78 |
val this_prefix = make_prefix context; |
|
79 |
in PrettyBlock (ind, consistent, context, map (add_prefix this_prefix) prts) end |
|
80 |
| add_prefix prefix (prt as PrettyString s) = |
|
81 |
if prefix = "" then prt else PrettyString (prefix ^ s) |
|
82 |
| add_prefix _ (prt as PrettyBreak _) = prt; |
|
83 |
||
84 |
val depth = get_print_depth (); |
|
85 |
val with_struct = ! PolyML.Compiler.printTypesWithStructureName; |
|
86 |
||
87 |
fun display disp x = |
|
88 |
if depth > 0 then |
|
89 |
(disp x |
|
90 |
|> with_struct ? add_prefix "" |
|
91 |
|> pretty_ml |> Pretty.from_ML |> Pretty.string_of |> put; put "\n") |
|
92 |
else (); |
|
93 |
||
94 |
fun apply_fix (a, b) = |
|
95 |
(display PolyML.NameSpace.displayFix (a, b); #enterFix space (a, b)); |
|
96 |
fun apply_type (a, b) = |
|
97 |
(display PolyML.NameSpace.displayType (b, depth); #enterType space (a, b)); |
|
98 |
fun apply_sig (a, b) = |
|
99 |
(display PolyML.NameSpace.displaySig (b, depth, space); #enterSig space (a, b)); |
|
100 |
fun apply_struct (a, b) = |
|
101 |
(display PolyML.NameSpace.displayStruct (b, depth, space); #enterStruct space (a, b)); |
|
102 |
fun apply_funct (a, b) = |
|
103 |
(display PolyML.NameSpace.displayFunct (b, depth, space); #enterFunct space (a, b)); |
|
104 |
fun apply_val (a, b) = |
|
105 |
(display PolyML.NameSpace.displayVal (b, depth, space); #enterVal space (a, b)); |
|
106 |
in |
|
107 |
List.app apply_fix fixes; |
|
108 |
List.app apply_type types; |
|
109 |
List.app apply_sig signatures; |
|
110 |
List.app apply_struct structures; |
|
111 |
List.app apply_funct functors; |
|
112 |
List.app apply_val values |
|
113 |
end; |
|
114 |
||
| 30646 | 115 |
fun result_fun (parse_tree, code) () = |
| 30668 | 116 |
(Context.>> (Result.map (append (the_list parse_tree))); |
| 30681 | 117 |
(case code of NONE => error "Static Errors" | SOME result => apply_result (result ()))); |
| 30680 | 118 |
|
119 |
||
120 |
(* compiler invocation *) |
|
| 30646 | 121 |
|
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
122 |
val parameters = |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
123 |
[PolyML.Compiler.CPOutStream put, |
| 30680 | 124 |
PolyML.Compiler.CPNameSpace space, |
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
125 |
PolyML.Compiler.CPErrorMessageProc put_message, |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
126 |
PolyML.Compiler.CPLineNo (fn () => ! current_line), |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
127 |
PolyML.Compiler.CPLineOffset get_index, |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
128 |
PolyML.Compiler.CPFileName (the_default "ML" (Position.file_of pos)), |
| 30680 | 129 |
PolyML.Compiler.CPPrintInAlphabeticalOrder false, |
130 |
PolyML.Compiler.CPCompilerResultFun result_fun]; |
|
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
131 |
val _ = |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
132 |
(while not (List.null (! in_buffer)) do |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
133 |
PolyML.compiler (get, parameters) ()) |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
134 |
handle exn => |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
135 |
(put ("Exception- " ^ General.exnMessage exn ^ " raised");
|
| 30677 | 136 |
error (output ()); raise exn); |
| 30680 | 137 |
in if verbose then print (output ()) else () end; |
138 |
||
139 |
val eval = use_text ML_Context.local_context; |
|
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
140 |
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
141 |
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
142 |
(* ML test command *) |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
143 |
|
| 30680 | 144 |
fun ML_test (txt, pos) = |
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
145 |
let |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
146 |
val _ = Position.report Markup.ML_source pos; |
| 30644 | 147 |
val ants = ML_Lex.read_antiq (Symbol_Pos.explode (txt, pos), pos); |
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
148 |
val ((env, body), env_ctxt) = ML_Context.eval_antiquotes (ants, pos) (Context.thread_data ()); |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
149 |
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
150 |
val _ = Context.setmp_thread_data env_ctxt |
| 30680 | 151 |
(fn () => (eval false Position.none env; Context.thread_data ())) () |
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
152 |
|> (fn NONE => () | SOME context' => Context.>> (ML_Context.inherit_env context')); |
| 30680 | 153 |
val _ = eval true pos body; |
154 |
val _ = eval false Position.none (ML_Lex.tokenize "structure Isabelle = struct end"); |
|
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
155 |
in () end; |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
156 |
|
| 30644 | 157 |
|
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
158 |
local structure P = OuterParse and K = OuterKeyword in |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
159 |
|
| 30644 | 160 |
fun inherit_env (context as Context.Proof lthy) = |
161 |
Context.Proof (LocalTheory.map_contexts (ML_Context.inherit_env context) lthy) |
|
162 |
| inherit_env context = context; |
|
163 |
||
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
164 |
val _ = |
| 30644 | 165 |
OuterSyntax.command "ML_test" "advanced ML compiler test" (K.tag_ml K.thy_decl) |
166 |
(P.ML_source >> (fn src => |
|
| 30680 | 167 |
Toplevel.generic_theory (ML_Context.exec (fn () => ML_test src) #> inherit_env))); |
|
30640
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
168 |
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
169 |
end; |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
170 |
|
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
171 |
end; |
|
3f3d1e218b64
Test of advanced ML compiler invocation in Poly/ML 5.3.
wenzelm
parents:
diff
changeset
|
172 |