| author | haftmann | 
| Wed, 23 Dec 2009 08:31:15 +0100 | |
| changeset 34173 | 458ced35abb8 | 
| parent 33603 | 3713a5208671 | 
| child 37209 | 1c8cf0048934 | 
| permissions | -rw-r--r-- | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/ML/ml_compiler_polyml-5.3.ML  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 
33538
 
edf497b5b5d2
setup for official Poly/ML 5.3.0, which is now the default;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
4  | 
Advanced runtime compilation for Poly/ML 5.3.0.  | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
signature ML_COMPILER =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 31477 | 9  | 
val exn_position: exn -> Position.T  | 
10  | 
val exn_message: exn -> string  | 
|
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
val eval: bool -> Position.T -> ML_Lex.token list -> unit  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
end  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
structure ML_Compiler: ML_COMPILER =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
15  | 
struct  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
|
| 
31429
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
17  | 
(* source locations *)  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
18  | 
|
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
19  | 
fun position_of (loc: PolyML.location) =  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
20  | 
let  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
21  | 
    val {file = text, startLine = line, startPosition = offset,
 | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
22  | 
endLine = end_line, endPosition = end_offset} = loc;  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
23  | 
val loc_props =  | 
| 
31434
 
1f9b6a5dc8cc
more robust treatment of bootstrap source positions;
 
wenzelm 
parents: 
31429 
diff
changeset
 | 
24  | 
(case YXML.parse text of  | 
| 
 
1f9b6a5dc8cc
more robust treatment of bootstrap source positions;
 
wenzelm 
parents: 
31429 
diff
changeset
 | 
25  | 
XML.Elem (e, atts, _) => if e = Markup.positionN then atts else []  | 
| 
 
1f9b6a5dc8cc
more robust treatment of bootstrap source positions;
 
wenzelm 
parents: 
31429 
diff
changeset
 | 
26  | 
| XML.Text s => Position.file_name s);  | 
| 
31429
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
27  | 
in  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
28  | 
Position.value Markup.lineN line @  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
29  | 
Position.value Markup.offsetN offset @  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
30  | 
Position.value Markup.end_lineN end_line @  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
31  | 
Position.value Markup.end_offsetN end_offset @  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
32  | 
loc_props  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
33  | 
end |> Position.of_properties;  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
34  | 
|
| 31477 | 35  | 
fun exn_position exn =  | 
| 
31429
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
36  | 
(case PolyML.exceptionLocation exn of  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
37  | 
NONE => Position.none  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
38  | 
| SOME loc => position_of loc);  | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
39  | 
|
| 31477 | 40  | 
val exn_message = Runtime.exn_message exn_position;  | 
41  | 
||
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
42  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
(* parse trees *)  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
|
| 
31429
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
45  | 
fun report_parse_tree depth space =  | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
let  | 
| 31475 | 47  | 
fun report_decl markup loc decl =  | 
48  | 
Position.report_text Markup.ML_ref (position_of loc)  | 
|
49  | 
(Markup.markup (Markup.properties (Position.properties_of (position_of decl)) markup) "");  | 
|
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
fun report loc (PolyML.PTtype types) =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
PolyML.NameSpace.displayTypeExpression (types, depth, space)  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
|> pretty_ml |> Pretty.from_ML |> Pretty.string_of  | 
| 
31429
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
53  | 
|> Position.report_text Markup.ML_typing (position_of loc)  | 
| 31475 | 54  | 
| report loc (PolyML.PTdeclaredAt decl) = report_decl Markup.ML_def loc decl  | 
55  | 
| report loc (PolyML.PTopenedAt decl) = report_decl Markup.ML_open loc decl  | 
|
56  | 
| report loc (PolyML.PTstructureAt decl) = report_decl Markup.ML_struct loc decl  | 
|
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
| report _ (PolyML.PTnextSibling tree) = report_tree (tree ())  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
58  | 
| report _ (PolyML.PTfirstChild tree) = report_tree (tree ())  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
| report _ _ = ()  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
and report_tree (loc, props) = List.app (report loc) props;  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
in report_tree end;  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
(* eval ML source tokens *)  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
65  | 
|
| 
31756
 
178621145f98
eval: tuned input buffer and get_offset (which ran very hot, up to the point where AFP/FlyspeckTame would seemingly take forever);
 
wenzelm 
parents: 
31495 
diff
changeset
 | 
66  | 
val offset_of = the_default 0 o Position.offset_of;  | 
| 
 
178621145f98
eval: tuned input buffer and get_offset (which ran very hot, up to the point where AFP/FlyspeckTame would seemingly take forever);
 
wenzelm 
parents: 
31495 
diff
changeset
 | 
67  | 
|
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
68  | 
fun eval verbose pos toks =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
69  | 
let  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
val _ = Secure.secure_mltext ();  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
71  | 
    val {name_space = space, print, error = err, ...} = ML_Env.local_context;
 | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
74  | 
(* input *)  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
75  | 
|
| 31437 | 76  | 
val location_props =  | 
| 
32492
 
9d49a280f3f9
eval/location_props: always produce YXML markup, independent of print_mode;
 
wenzelm 
parents: 
32490 
diff
changeset
 | 
77  | 
op ^ (YXML.output_markup (Markup.position |> Markup.properties  | 
| 
 
9d49a280f3f9
eval/location_props: always produce YXML markup, independent of print_mode;
 
wenzelm 
parents: 
32490 
diff
changeset
 | 
78  | 
(filter (member (op =) [Markup.idN, Markup.fileN] o #1) (Position.properties_of pos))));  | 
| 31437 | 79  | 
|
| 
31429
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
80  | 
val input = toks |> maps (fn tok =>  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
81  | 
let  | 
| 
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
82  | 
val syms = Symbol.explode (ML_Lex.check_content_of tok);  | 
| 31437 | 83  | 
val (ps, _) = fold_map (fn s => fn p => (p, Position.advance s p)) syms  | 
84  | 
(Position.reset_range (ML_Lex.pos_of tok));  | 
|
| 
31756
 
178621145f98
eval: tuned input buffer and get_offset (which ran very hot, up to the point where AFP/FlyspeckTame would seemingly take forever);
 
wenzelm 
parents: 
31495 
diff
changeset
 | 
85  | 
in  | 
| 
 
178621145f98
eval: tuned input buffer and get_offset (which ran very hot, up to the point where AFP/FlyspeckTame would seemingly take forever);
 
wenzelm 
parents: 
31495 
diff
changeset
 | 
86  | 
(ps ~~ syms) |> maps (fn (p, sym) =>  | 
| 
 
178621145f98
eval: tuned input buffer and get_offset (which ran very hot, up to the point where AFP/FlyspeckTame would seemingly take forever);
 
wenzelm 
parents: 
31495 
diff
changeset
 | 
87  | 
map (pair (offset_of p)) (String.explode (Symbol.esc sym)))  | 
| 
 
178621145f98
eval: tuned input buffer and get_offset (which ran very hot, up to the point where AFP/FlyspeckTame would seemingly take forever);
 
wenzelm 
parents: 
31495 
diff
changeset
 | 
88  | 
end);  | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
89  | 
|
| 31475 | 90  | 
val end_pos =  | 
91  | 
if null toks then Position.none  | 
|
92  | 
else ML_Lex.end_pos_of (List.last toks);  | 
|
93  | 
||
| 32738 | 94  | 
val input_buffer = Unsynchronized.ref (input @ [(offset_of end_pos, #"\n")]);  | 
95  | 
val line = Unsynchronized.ref (the_default 1 (Position.line_of pos));  | 
|
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
96  | 
|
| 
31756
 
178621145f98
eval: tuned input buffer and get_offset (which ran very hot, up to the point where AFP/FlyspeckTame would seemingly take forever);
 
wenzelm 
parents: 
31495 
diff
changeset
 | 
97  | 
fun get_offset () = (case ! input_buffer of [] => 0 | (i, _) :: _ => i);  | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
fun get () =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
(case ! input_buffer of  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
100  | 
[] => NONE  | 
| 
31756
 
178621145f98
eval: tuned input buffer and get_offset (which ran very hot, up to the point where AFP/FlyspeckTame would seemingly take forever);
 
wenzelm 
parents: 
31495 
diff
changeset
 | 
101  | 
| (_, c) :: rest =>  | 
| 
 
178621145f98
eval: tuned input buffer and get_offset (which ran very hot, up to the point where AFP/FlyspeckTame would seemingly take forever);
 
wenzelm 
parents: 
31495 
diff
changeset
 | 
102  | 
(input_buffer := rest;  | 
| 
31429
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
103  | 
if c = #"\n" then line := ! line + 1 else ();  | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
SOME c));  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
107  | 
(* output *)  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
|
| 32738 | 109  | 
val output_buffer = Unsynchronized.ref Buffer.empty;  | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
110  | 
fun output () = Buffer.content (! output_buffer);  | 
| 32738 | 111  | 
fun put s = Unsynchronized.change output_buffer (Buffer.add s);  | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
112  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
113  | 
    fun put_message {message, hard, location, context = _} =
 | 
| 32490 | 114  | 
(put ((if hard then "Error" else "Warning") ^ Position.str_of (position_of location) ^ ":\n");  | 
115  | 
put (Pretty.string_of (Pretty.from_ML (pretty_ml message)) ^ "\n"));  | 
|
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
117  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
118  | 
(* results *)  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
119  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
120  | 
val depth = get_print_depth ();  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
121  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
122  | 
    fun apply_result {fixes, types, signatures, structures, functors, values} =
 | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
123  | 
let  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
124  | 
fun display disp x =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
125  | 
if depth > 0 then  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
126  | 
(disp x |> pretty_ml |> Pretty.from_ML |> Pretty.string_of |> put; put "\n")  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
127  | 
else ();  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
128  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
fun apply_fix (a, b) =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
130  | 
(display PolyML.NameSpace.displayFix (a, b); #enterFix space (a, b));  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
131  | 
fun apply_type (a, b) =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
132  | 
(display PolyML.NameSpace.displayType (b, depth, space); #enterType space (a, b));  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
133  | 
fun apply_sig (a, b) =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
134  | 
(display PolyML.NameSpace.displaySig (b, depth, space); #enterSig space (a, b));  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
135  | 
fun apply_struct (a, b) =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
(display PolyML.NameSpace.displayStruct (b, depth, space); #enterStruct space (a, b));  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
137  | 
fun apply_funct (a, b) =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
138  | 
(display PolyML.NameSpace.displayFunct (b, depth, space); #enterFunct space (a, b));  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
139  | 
fun apply_val (a, b) =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
140  | 
(display PolyML.NameSpace.displayVal (b, depth, space); #enterVal space (a, b));  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
141  | 
in  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
142  | 
List.app apply_fix fixes;  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
143  | 
List.app apply_type types;  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
List.app apply_sig signatures;  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
145  | 
List.app apply_struct structures;  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
146  | 
List.app apply_funct functors;  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
147  | 
List.app apply_val values  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
148  | 
end;  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
150  | 
fun result_fun (phase1, phase2) () =  | 
| 31477 | 151  | 
((case phase1 of  | 
152  | 
NONE => ()  | 
|
153  | 
| SOME parse_tree => report_parse_tree depth space parse_tree);  | 
|
154  | 
(case phase2 of  | 
|
155  | 
NONE => err "Static Errors"  | 
|
156  | 
| SOME code =>  | 
|
| 
33603
 
3713a5208671
generalized Runtime.toplevel_error wrt. output function;
 
wenzelm 
parents: 
33538 
diff
changeset
 | 
157  | 
apply_result  | 
| 
 
3713a5208671
generalized Runtime.toplevel_error wrt. output function;
 
wenzelm 
parents: 
33538 
diff
changeset
 | 
158  | 
((code  | 
| 
 
3713a5208671
generalized Runtime.toplevel_error wrt. output function;
 
wenzelm 
parents: 
33538 
diff
changeset
 | 
159  | 
|> Runtime.debugging  | 
| 
 
3713a5208671
generalized Runtime.toplevel_error wrt. output function;
 
wenzelm 
parents: 
33538 
diff
changeset
 | 
160  | 
|> Runtime.toplevel_error (Output.error_msg o exn_message)) ())));  | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
161  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
162  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
163  | 
(* compiler invocation *)  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
164  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
165  | 
val parameters =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
166  | 
[PolyML.Compiler.CPOutStream put,  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
167  | 
PolyML.Compiler.CPNameSpace space,  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
168  | 
PolyML.Compiler.CPErrorMessageProc put_message,  | 
| 
31429
 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 
wenzelm 
parents: 
31333 
diff
changeset
 | 
169  | 
PolyML.Compiler.CPLineNo (fn () => ! line),  | 
| 31475 | 170  | 
PolyML.Compiler.CPLineOffset get_offset,  | 
| 31437 | 171  | 
PolyML.Compiler.CPFileName location_props,  | 
| 31475 | 172  | 
PolyML.Compiler.CPCompilerResultFun result_fun,  | 
173  | 
PolyML.Compiler.CPPrintInAlphabeticalOrder false];  | 
|
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
174  | 
val _ =  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
175  | 
(while not (List.null (! input_buffer)) do  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
176  | 
PolyML.compiler (get, parameters) ())  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
177  | 
handle exn =>  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
178  | 
       (put ("Exception- " ^ General.exnMessage exn ^ " raised");
 | 
| 
31480
 
05937d6aafb5
reraise exceptions to preserve position information;
 
wenzelm 
parents: 
31477 
diff
changeset
 | 
179  | 
err (output ()); reraise exn);  | 
| 
31333
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
180  | 
in if verbose then print (output ()) else () end;  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
181  | 
|
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
182  | 
end;  | 
| 
 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 
wenzelm 
parents:  
diff
changeset
 | 
183  |