--- a/src/Pure/ML/ml_compiler_polyml.ML Fri Oct 12 12:21:01 2012 +0200
+++ b/src/Pure/ML/ml_compiler_polyml.ML Fri Oct 12 14:05:30 2012 +0200
@@ -113,10 +113,8 @@
let
val pos = position_of loc;
val txt =
- (Position.is_reported pos ? Markup.markup Isabelle_Markup.no_report)
- ((if hard then "Error" else "Warning") ^ Position.here pos ^ ":\n") ^
- Pretty.string_of (Pretty.from_ML (pretty_ml msg)) ^
- Position.reported_text pos Isabelle_Markup.report "";
+ (if hard then "ML error" else "ML warning") ^ Position.here pos ^ ":\n" ^
+ Pretty.string_of (Pretty.from_ML (pretty_ml msg));
in if hard then err txt else warn txt end;
--- a/src/Pure/goal.ML Fri Oct 12 12:21:01 2012 +0200
+++ b/src/Pure/goal.ML Fri Oct 12 14:05:30 2012 +0200
@@ -160,9 +160,11 @@
(case result of
Exn.Res _ => ()
| Exn.Exn exn =>
- (status task [Isabelle_Markup.failed];
- Output.report (Markup.markup_only Isabelle_Markup.bad);
- Output.error_msg (ML_Compiler.exn_message exn)));
+ if id = 0 orelse Exn.is_interrupt exn then ()
+ else
+ (status task [Isabelle_Markup.failed];
+ Output.report (Markup.markup_only Isabelle_Markup.bad);
+ Output.error_msg (ML_Compiler.exn_message exn)));
val _ = count_forked ~1;
in Exn.release result end);
val _ = status (Future.task_of future) [Isabelle_Markup.forked];