1 (* Title: Pure/ML/ml_test.ML |
|
2 Author: Makarius |
|
3 |
|
4 Test of advanced ML compiler invocation in Poly/ML 5.3 (SVN 744). |
|
5 *) |
|
6 |
|
7 signature ML_TEST = |
|
8 sig |
|
9 val position_of: Proof.context -> PolyML.location -> Position.T |
|
10 val eval: bool -> Position.T -> ML_Lex.token list -> unit |
|
11 end |
|
12 |
|
13 structure ML_Test: ML_TEST = |
|
14 struct |
|
15 |
|
16 (* extra ML environment *) |
|
17 |
|
18 structure Extra_Env = GenericDataFun |
|
19 ( |
|
20 type T = Position.T Inttab.table; (*position of registered tokens*) |
|
21 val empty = Inttab.empty; |
|
22 val extend = I; |
|
23 fun merge _ = Inttab.merge (K true); |
|
24 ); |
|
25 |
|
26 fun inherit_env context = |
|
27 ML_Context.inherit_env context #> |
|
28 Extra_Env.put (Extra_Env.get context); |
|
29 |
|
30 fun register_tokens toks context = |
|
31 let |
|
32 val regs = map (fn tok => (serial (), tok)) toks; |
|
33 val context' = context |
|
34 |> Extra_Env.map (fold (fn (i, tok) => Inttab.update (i, ML_Lex.pos_of tok)) regs); |
|
35 in (regs, context') end; |
|
36 |
|
37 fun position_of ctxt |
|
38 ({file, startLine = line, startPosition = i, endPosition = j, ...}: PolyML.location) = |
|
39 (case pairself (Inttab.lookup (Extra_Env.get (Context.Proof ctxt))) (i, j) of |
|
40 (SOME pos1, SOME pos2) => Position.encode_range (pos1, pos2) |
|
41 | (SOME pos, NONE) => pos |
|
42 | _ => Position.line_file line file); |
|
43 |
|
44 |
|
45 (* parse trees *) |
|
46 |
|
47 fun report_parse_tree context depth space = |
|
48 let |
|
49 val pos_of = position_of (Context.proof_of context); |
|
50 fun report loc (PolyML.PTtype types) = |
|
51 PolyML.NameSpace.displayTypeExpression (types, depth, space) |
|
52 |> pretty_ml |> Pretty.from_ML |> Pretty.string_of |
|
53 |> Position.report_text Markup.ML_typing (pos_of loc) |
|
54 | report loc (PolyML.PTdeclaredAt decl) = |
|
55 Markup.markup (Markup.properties (Position.properties_of (pos_of decl)) Markup.ML_def) "" |
|
56 |> Position.report_text Markup.ML_ref (pos_of loc) |
|
57 | report _ (PolyML.PTnextSibling tree) = report_tree (tree ()) |
|
58 | report _ (PolyML.PTfirstChild tree) = report_tree (tree ()) |
|
59 | report _ _ = () |
|
60 and report_tree (loc, props) = List.app (report loc) props; |
|
61 in report_tree end; |
|
62 |
|
63 |
|
64 (* eval ML source tokens *) |
|
65 |
|
66 fun use_text ({name_space = space, print, error, ...}: use_context) verbose pos toks = |
|
67 let |
|
68 (* input *) |
|
69 |
|
70 val input = Context.>>> (register_tokens toks); |
|
71 val input_buffer = ref (map (apsnd (String.explode o ML_Lex.text_of)) input); |
|
72 |
|
73 val current_line = ref (the_default 1 (Position.line_of pos)); |
|
74 |
|
75 fun get_index () = |
|
76 the_default 0 (get_first (fn (_, []) => NONE | (i, _) => SOME i) (! input_buffer)); |
|
77 |
|
78 fun get () = |
|
79 (case ! input_buffer of |
|
80 [] => NONE |
|
81 | (_, []) :: rest => (input_buffer := rest; get ()) |
|
82 | (i, c :: cs) :: rest => |
|
83 (input_buffer := (i, cs) :: rest; |
|
84 if c = #"\n" then current_line := ! current_line + 1 else (); |
|
85 SOME c)); |
|
86 |
|
87 |
|
88 (* output *) |
|
89 |
|
90 val output_buffer = ref Buffer.empty; |
|
91 fun output () = Buffer.content (! output_buffer); |
|
92 fun put s = change output_buffer (Buffer.add s); |
|
93 |
|
94 fun put_message {message, hard, location, context = _} = |
|
95 (put (if hard then "Error: " else "Warning: "); |
|
96 put (Pretty.string_of (Pretty.from_ML (pretty_ml message))); |
|
97 put (Position.str_of |
|
98 (position_of (Context.proof_of (the (Context.thread_data ()))) location) ^ "\n")); |
|
99 |
|
100 |
|
101 (* results *) |
|
102 |
|
103 val depth = get_print_depth (); |
|
104 |
|
105 fun apply_result {fixes, types, signatures, structures, functors, values} = |
|
106 let |
|
107 fun display disp x = |
|
108 if depth > 0 then |
|
109 (disp x |> pretty_ml |> Pretty.from_ML |> Pretty.string_of |> put; put "\n") |
|
110 else (); |
|
111 |
|
112 fun apply_fix (a, b) = |
|
113 (display PolyML.NameSpace.displayFix (a, b); #enterFix space (a, b)); |
|
114 fun apply_type (a, b) = |
|
115 (display PolyML.NameSpace.displayType (b, depth, space); #enterType space (a, b)); |
|
116 fun apply_sig (a, b) = |
|
117 (display PolyML.NameSpace.displaySig (b, depth, space); #enterSig space (a, b)); |
|
118 fun apply_struct (a, b) = |
|
119 (display PolyML.NameSpace.displayStruct (b, depth, space); #enterStruct space (a, b)); |
|
120 fun apply_funct (a, b) = |
|
121 (display PolyML.NameSpace.displayFunct (b, depth, space); #enterFunct space (a, b)); |
|
122 fun apply_val (a, b) = |
|
123 (display PolyML.NameSpace.displayVal (b, depth, space); #enterVal space (a, b)); |
|
124 in |
|
125 List.app apply_fix fixes; |
|
126 List.app apply_type types; |
|
127 List.app apply_sig signatures; |
|
128 List.app apply_struct structures; |
|
129 List.app apply_funct functors; |
|
130 List.app apply_val values |
|
131 end; |
|
132 |
|
133 fun result_fun (phase1, phase2) () = |
|
134 (case phase1 of NONE => () |
|
135 | SOME parse_tree => report_parse_tree (the (Context.thread_data ())) depth space parse_tree; |
|
136 case phase2 of NONE => error "Static Errors" |
|
137 | SOME code => apply_result (Toplevel.program code)); |
|
138 |
|
139 |
|
140 (* compiler invocation *) |
|
141 |
|
142 val parameters = |
|
143 [PolyML.Compiler.CPOutStream put, |
|
144 PolyML.Compiler.CPNameSpace space, |
|
145 PolyML.Compiler.CPErrorMessageProc put_message, |
|
146 PolyML.Compiler.CPLineNo (fn () => ! current_line), |
|
147 PolyML.Compiler.CPFileName (the_default "ML" (Position.file_of pos)), |
|
148 PolyML.Compiler.CPLineOffset get_index, |
|
149 PolyML.Compiler.CPCompilerResultFun result_fun]; |
|
150 val _ = |
|
151 (while not (List.null (! input_buffer)) do |
|
152 PolyML.compiler (get, parameters) ()) |
|
153 handle exn => |
|
154 (put ("Exception- " ^ General.exnMessage exn ^ " raised"); |
|
155 error (output ()); raise exn); |
|
156 in if verbose then print (output ()) else () end; |
|
157 |
|
158 val eval = use_text ML_Context.local_context; |
|
159 |
|
160 |
|
161 (* ML test command *) |
|
162 |
|
163 fun ML_test (txt, pos) = |
|
164 let |
|
165 val _ = Position.report Markup.ML_source pos; |
|
166 val ants = ML_Lex.read_antiq (Symbol_Pos.explode (txt, pos), pos); |
|
167 val ((env, body), env_ctxt) = ML_Context.eval_antiquotes (ants, pos) (Context.thread_data ()); |
|
168 |
|
169 val _ = Context.setmp_thread_data env_ctxt |
|
170 (fn () => (eval false Position.none env; Context.thread_data ())) () |
|
171 |> (fn NONE => () | SOME context' => Context.>> (inherit_env context')); |
|
172 val _ = eval true pos body; |
|
173 val _ = eval false Position.none (ML_Lex.tokenize "structure Isabelle = struct end"); |
|
174 in () end; |
|
175 |
|
176 |
|
177 local structure P = OuterParse and K = OuterKeyword in |
|
178 |
|
179 fun propagate_env (context as Context.Proof lthy) = |
|
180 Context.Proof (LocalTheory.map_contexts (inherit_env context) lthy) |
|
181 | propagate_env context = context; |
|
182 |
|
183 val _ = |
|
184 OuterSyntax.command "ML_test" "advanced ML compiler test" (K.tag_ml K.thy_decl) |
|
185 (P.ML_source >> (fn src => |
|
186 Toplevel.generic_theory (ML_Context.exec (fn () => ML_test src) #> propagate_env))); |
|
187 |
|
188 end; |
|
189 |
|
190 end; |
|
191 |
|