| author | kleing | 
| Thu, 20 Oct 2011 12:30:43 -0400 | |
| changeset 45218 | f115540543d8 | 
| parent 45147 | c23029f6357f | 
| child 45666 | d83797ef0d2d | 
| 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 | |
| 45147 | 4 | Advanced runtime compilation for Poly/ML 5.3.0 or later. | 
| 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 | 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 | 8 | struct | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 9 | |
| 31429 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 10 | (* source locations *) | 
| 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 11 | |
| 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 12 | fun position_of (loc: PolyML.location) = | 
| 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 13 | let | 
| 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 14 |     val {file = text, startLine = line, startPosition = offset,
 | 
| 41483 | 15 | endLine = _, endPosition = end_offset} = loc; | 
| 41484 | 16 | val props = | 
| 31434 
1f9b6a5dc8cc
more robust treatment of bootstrap source positions;
 wenzelm parents: 
31429diff
changeset | 17 | (case YXML.parse text of | 
| 38228 
ada3ab6b9085
simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
 wenzelm parents: 
37209diff
changeset | 18 | XML.Elem ((e, atts), _) => if e = Markup.positionN then atts else [] | 
| 31434 
1f9b6a5dc8cc
more robust treatment of bootstrap source positions;
 wenzelm parents: 
31429diff
changeset | 19 | | XML.Text s => Position.file_name s); | 
| 31429 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 20 | in | 
| 43710 
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
 wenzelm parents: 
42327diff
changeset | 21 |     Position.make {line = line, offset = offset, end_offset = end_offset, props = props}
 | 
| 41484 | 22 | end; | 
| 31429 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 23 | |
| 31477 | 24 | fun exn_position exn = | 
| 31429 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 25 | (case PolyML.exceptionLocation exn of | 
| 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 26 | NONE => Position.none | 
| 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 27 | | 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 | 28 | |
| 38874 
4a4d34d2f97b
more careful treatment of nested multi-exceptions, collapsing cumulatively empty list to Interrupt;
 wenzelm parents: 
38720diff
changeset | 29 | val exn_messages = Runtime.exn_messages exn_position; | 
| 31477 | 30 | val exn_message = Runtime.exn_message exn_position; | 
| 31 | ||
| 31333 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 32 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 33 | (* parse trees *) | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 34 | |
| 31429 
8a7c0899e0b1
convert explicitly between Position.T/PolyML.location, without costly registration of tokens;
 wenzelm parents: 
31333diff
changeset | 35 | 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 | 36 | let | 
| 44737 | 37 | val reported_text = | 
| 38720 
7f8bc335e203
ML_Context.eval: produce antiquotation environment preferably in invisible context, to avoid displaced report messages from ML_Compiler;
 wenzelm parents: 
38228diff
changeset | 38 | (case Context.thread_data () of | 
| 44737 | 39 | SOME (Context.Proof ctxt) => Context_Position.reported_text ctxt | 
| 40 | | _ => Position.reported_text); | |
| 38720 
7f8bc335e203
ML_Context.eval: produce antiquotation environment preferably in invisible context, to avoid displaced report messages from ML_Compiler;
 wenzelm parents: 
38228diff
changeset | 41 | |
| 44737 | 42 | fun reported_entity kind loc decl = | 
| 43 | reported_text (position_of loc) | |
| 42327 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41504diff
changeset | 44 | (Markup.entityN, (Markup.kindN, kind) :: Position.def_properties_of (position_of decl)) ""; | 
| 41503 
a7462e442e35
refined report_parse_tree: reverse reports happen to produce proper type information for inlined @{term}, @{typ} etc.;
 wenzelm parents: 
41501diff
changeset | 45 | |
| 44737 | 46 | fun reported loc (PolyML.PTtype types) = | 
| 47 | cons | |
| 48 | (PolyML.NameSpace.displayTypeExpression (types, depth, space) | |
| 49 | |> pretty_ml |> Pretty.from_ML |> Pretty.string_of | |
| 50 | |> reported_text (position_of loc) Markup.ML_typing) | |
| 51 | | reported loc (PolyML.PTdeclaredAt decl) = | |
| 52 | cons (reported_entity Markup.ML_defN loc decl) | |
| 53 | | reported loc (PolyML.PTopenedAt decl) = | |
| 54 | cons (reported_entity Markup.ML_openN loc decl) | |
| 55 | | reported loc (PolyML.PTstructureAt decl) = | |
| 56 | cons (reported_entity Markup.ML_structN loc decl) | |
| 57 | | reported _ (PolyML.PTnextSibling tree) = reported_tree (tree ()) | |
| 58 | | reported _ (PolyML.PTfirstChild tree) = reported_tree (tree ()) | |
| 59 | | reported _ _ = I | |
| 60 | and reported_tree (loc, props) = fold (reported loc) props; | |
| 61 | in fn tree => Output.report (implode (reported_tree tree [])) end; | |
| 31333 
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 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 66 | 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 | 67 | let | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 68 | val _ = Secure.secure_mltext (); | 
| 39227 | 69 | val space = ML_Env.name_space; | 
| 31333 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 70 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 71 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 72 | (* input *) | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 73 | |
| 31437 | 74 | val location_props = | 
| 32492 
9d49a280f3f9
eval/location_props: always produce YXML markup, independent of print_mode;
 wenzelm parents: 
32490diff
changeset | 75 | op ^ (YXML.output_markup (Markup.position |> Markup.properties | 
| 
9d49a280f3f9
eval/location_props: always produce YXML markup, independent of print_mode;
 wenzelm parents: 
32490diff
changeset | 76 | (filter (member (op =) [Markup.idN, Markup.fileN] o #1) (Position.properties_of pos)))); | 
| 31437 | 77 | |
| 41501 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 78 | val input_buffer = | 
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 79 | Unsynchronized.ref (toks |> map | 
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 80 | (`(maps (String.explode o Symbol.esc) o Symbol.explode o ML_Lex.check_content_of))); | 
| 31333 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 81 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 82 | fun get () = | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 83 | (case ! input_buffer of | 
| 41501 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 84 | (c :: cs, tok) :: rest => (input_buffer := (cs, tok) :: rest; SOME c) | 
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 85 | | ([], _) :: rest => (input_buffer := rest; SOME #" ") | 
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 86 | | [] => NONE); | 
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 87 | |
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 88 | fun get_pos () = | 
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 89 | (case ! input_buffer of | 
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 90 | (_ :: _, tok) :: _ => ML_Lex.pos_of tok | 
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 91 | | ([], tok) :: _ => ML_Lex.end_pos_of tok | 
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 92 | | [] => Position.none); | 
| 31333 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 93 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 94 | |
| 39228 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 95 | (* output channels *) | 
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 96 | |
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 97 | val writeln_buffer = Unsynchronized.ref Buffer.empty; | 
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 98 | fun write s = Unsynchronized.change writeln_buffer (Buffer.add s); | 
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 99 | fun output_writeln () = writeln (Buffer.content (! writeln_buffer)); | 
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 100 | |
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 101 | val warnings = Unsynchronized.ref ([]: string list); | 
| 39231 
25c345302a17
avoid mixing of static and runtime errors in compiler output, to accomodate Proof General;
 wenzelm parents: 
39230diff
changeset | 102 | fun warn msg = Unsynchronized.change warnings (cons msg); | 
| 39228 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 103 | fun output_warnings () = List.app warning (rev (! warnings)); | 
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 104 | |
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 105 | val error_buffer = Unsynchronized.ref Buffer.empty; | 
| 39231 
25c345302a17
avoid mixing of static and runtime errors in compiler output, to accomodate Proof General;
 wenzelm parents: 
39230diff
changeset | 106 | fun err msg = Unsynchronized.change error_buffer (Buffer.add msg #> Buffer.add "\n"); | 
| 39228 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 107 | fun flush_error () = writeln (Buffer.content (! error_buffer)); | 
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 108 | fun raise_error msg = error (Buffer.content (Buffer.add msg (! error_buffer))); | 
| 31333 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 109 | |
| 39228 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 110 |     fun message {message = msg, hard, location = loc, context = _} =
 | 
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 111 | let | 
| 41504 | 112 | val pos = position_of loc; | 
| 39228 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 113 | val txt = | 
| 41504 | 114 | (Position.is_reported pos ? Markup.markup Markup.no_report) | 
| 115 | ((if hard then "Error" else "Warning") ^ Position.str_of pos ^ ":\n") ^ | |
| 39228 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 116 | Pretty.string_of (Pretty.from_ML (pretty_ml msg)) ^ | 
| 41504 | 117 | Position.reported_text pos Markup.report ""; | 
| 39231 
25c345302a17
avoid mixing of static and runtime errors in compiler output, to accomodate Proof General;
 wenzelm parents: 
39230diff
changeset | 118 | in if hard then err txt else warn txt end; | 
| 31333 
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 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 121 | (* results *) | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 122 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 123 | 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 | 124 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 125 |     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 | 126 | let | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 127 | 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 | 128 | if depth > 0 then | 
| 39228 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 129 | (disp x |> pretty_ml |> Pretty.from_ML |> Pretty.string_of |> write; write "\n") | 
| 31333 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 130 | else (); | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 131 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 132 | fun apply_fix (a, b) = | 
| 40799 
d44c87988a24
ML results: enter before printing (cf. Poly/ML SVN 1218);
 wenzelm parents: 
39507diff
changeset | 133 | (#enterFix space (a, b); display PolyML.NameSpace.displayFix (a, b)); | 
| 31333 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 134 | fun apply_type (a, b) = | 
| 40799 
d44c87988a24
ML results: enter before printing (cf. Poly/ML SVN 1218);
 wenzelm parents: 
39507diff
changeset | 135 | (#enterType space (a, b); display PolyML.NameSpace.displayType (b, 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 | 136 | fun apply_sig (a, b) = | 
| 40799 
d44c87988a24
ML results: enter before printing (cf. Poly/ML SVN 1218);
 wenzelm parents: 
39507diff
changeset | 137 | (#enterSig space (a, b); display PolyML.NameSpace.displaySig (b, 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 | 138 | fun apply_struct (a, b) = | 
| 40799 
d44c87988a24
ML results: enter before printing (cf. Poly/ML SVN 1218);
 wenzelm parents: 
39507diff
changeset | 139 | (#enterStruct space (a, b); display PolyML.NameSpace.displayStruct (b, 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 | 140 | fun apply_funct (a, b) = | 
| 40799 
d44c87988a24
ML results: enter before printing (cf. Poly/ML SVN 1218);
 wenzelm parents: 
39507diff
changeset | 141 | (#enterFunct space (a, b); display PolyML.NameSpace.displayFunct (b, 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 | 142 | fun apply_val (a, b) = | 
| 40799 
d44c87988a24
ML results: enter before printing (cf. Poly/ML SVN 1218);
 wenzelm parents: 
39507diff
changeset | 143 | (#enterVal space (a, b); display PolyML.NameSpace.displayVal (b, 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 | 144 | in | 
| 
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_fix fixes; | 
| 
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_type types; | 
| 
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_sig signatures; | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 148 | 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 | 149 | 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 | 150 | 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 | 151 | end; | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 152 | |
| 39231 
25c345302a17
avoid mixing of static and runtime errors in compiler output, to accomodate Proof General;
 wenzelm parents: 
39230diff
changeset | 153 | exception STATIC_ERRORS of unit; | 
| 39230 
184507f6e8d0
ML_Compiler.eval: discontinued extra "Static Errors" of raw Poly/ML;
 wenzelm parents: 
39228diff
changeset | 154 | |
| 31333 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 155 | fun result_fun (phase1, phase2) () = | 
| 31477 | 156 | ((case phase1 of | 
| 157 | NONE => () | |
| 158 | | SOME parse_tree => report_parse_tree depth space parse_tree); | |
| 159 | (case phase2 of | |
| 39231 
25c345302a17
avoid mixing of static and runtime errors in compiler output, to accomodate Proof General;
 wenzelm parents: 
39230diff
changeset | 160 | NONE => raise STATIC_ERRORS () | 
| 31477 | 161 | | SOME code => | 
| 33603 
3713a5208671
generalized Runtime.toplevel_error wrt. output function;
 wenzelm parents: 
33538diff
changeset | 162 | apply_result | 
| 
3713a5208671
generalized Runtime.toplevel_error wrt. output function;
 wenzelm parents: 
33538diff
changeset | 163 | ((code | 
| 
3713a5208671
generalized Runtime.toplevel_error wrt. output function;
 wenzelm parents: 
33538diff
changeset | 164 | |> Runtime.debugging | 
| 39231 
25c345302a17
avoid mixing of static and runtime errors in compiler output, to accomodate Proof General;
 wenzelm parents: 
39230diff
changeset | 165 | |> Runtime.toplevel_error (err 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 | 166 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 167 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 168 | (* compiler invocation *) | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 169 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 170 | val parameters = | 
| 39228 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 171 | [PolyML.Compiler.CPOutStream write, | 
| 31333 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 172 | PolyML.Compiler.CPNameSpace space, | 
| 39228 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 173 | PolyML.Compiler.CPErrorMessageProc message, | 
| 41501 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 174 | PolyML.Compiler.CPLineNo (the_default 0 o Position.line_of o get_pos), | 
| 
b5ad6b0d6d7c
ML compiler: more careful treatment of input tokens -- trailing space ensures proper separation and end position (cf. 82c1e348bc18, 08240feb69c7);
 wenzelm parents: 
41484diff
changeset | 175 | PolyML.Compiler.CPLineOffset (the_default 0 o Position.offset_of o get_pos), | 
| 31437 | 176 | PolyML.Compiler.CPFileName location_props, | 
| 31475 | 177 | PolyML.Compiler.CPCompilerResultFun result_fun, | 
| 178 | 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 | 179 | val _ = | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 180 | (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 | 181 | PolyML.compiler (get, parameters) ()) | 
| 39232 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 182 | handle exn => | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 183 | if Exn.is_interrupt exn then reraise exn | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 184 | else | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 185 | let | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 186 | val exn_msg = | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 187 | (case exn of | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 188 | STATIC_ERRORS () => "" | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 189 | | Runtime.TOPLEVEL_ERROR => "" | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 190 | | _ => "Exception- " ^ General.exnMessage exn ^ " raised"); | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 191 | val _ = output_warnings (); | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 192 | val _ = output_writeln (); | 
| 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 wenzelm parents: 
39231diff
changeset | 193 | in raise_error exn_msg end; | 
| 39228 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 194 | in | 
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 195 | if verbose then (output_warnings (); flush_error (); output_writeln ()) | 
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 196 | else () | 
| 
cb7264721c91
ML_Compiler.eval: more careful printing of messages and regular output, trying to accomodate Poly/ML, Proof General, Isabelle/Scala/jEdit at the same time;
 wenzelm parents: 
39227diff
changeset | 197 | end; | 
| 31333 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 198 | |
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 199 | end; | 
| 
fcd010617e6c
added structure ML_Compiler: runtime compilation, with advanced version for Poly/ML 5.3 (formerly ML_Test);
 wenzelm parents: diff
changeset | 200 |