src/Pure/System/isabelle_process.ML
changeset 49677 c4e2762a265c
parent 49661 ac48def96b69
child 50117 32755e357a51
--- a/src/Pure/System/isabelle_process.ML	Mon Oct 01 20:16:37 2012 +0200
+++ b/src/Pure/System/isabelle_process.ML	Mon Oct 01 20:17:30 2012 +0200
@@ -89,16 +89,16 @@
 
 fun chunk s = [string_of_int (size s), "\n", s];
 
-fun message do_flush mbox ch raw_props body =
+fun message do_flush mbox name raw_props body =
   let
     val robust_props = map (pairself YXML.embed_controls) raw_props;
-    val header = YXML.string_of (XML.Elem ((ch, robust_props), []));
+    val header = YXML.string_of (XML.Elem ((name, robust_props), []));
   in Mailbox.send mbox (chunk header @ chunk body, do_flush) end;
 
-fun standard_message mbox opt_serial ch body =
+fun standard_message mbox opt_serial name body =
   if body = "" then ()
   else
-    message false mbox ch
+    message false mbox name
       ((case opt_serial of SOME i => cons (Isabelle_Markup.serialN, string_of_int i) | _ => I)
         (Position.properties_of (Position.thread_data ()))) body;
 
@@ -124,17 +124,22 @@
     val mbox = Mailbox.create () : (string list * bool) Mailbox.T;
     val _ = Simple_Thread.fork false (message_output mbox channel);
   in
-    Output.Private_Hooks.status_fn := standard_message mbox NONE "B";
-    Output.Private_Hooks.report_fn := standard_message mbox NONE "C";
-    Output.Private_Hooks.writeln_fn := (fn s => standard_message mbox (SOME (serial ())) "D" s);
+    Output.Private_Hooks.status_fn := standard_message mbox NONE Isabelle_Markup.statusN;
+    Output.Private_Hooks.report_fn := standard_message mbox NONE Isabelle_Markup.reportN;
+    Output.Private_Hooks.writeln_fn :=
+      (fn s => standard_message mbox (SOME (serial ())) Isabelle_Markup.writelnN s);
     Output.Private_Hooks.tracing_fn :=
-      (fn s => (update_tracing_limits s; standard_message mbox (SOME (serial ())) "E" s));
-    Output.Private_Hooks.warning_fn := (fn s => standard_message mbox (SOME (serial ())) "F" s);
-    Output.Private_Hooks.error_fn := (fn (i, s) => standard_message mbox (SOME i) "G" s);
-    Output.Private_Hooks.protocol_message_fn := message true mbox "H";
+      (fn s =>
+        (update_tracing_limits s;
+          standard_message mbox (SOME (serial ())) Isabelle_Markup.tracingN s));
+    Output.Private_Hooks.warning_fn :=
+      (fn s => standard_message mbox (SOME (serial ())) Isabelle_Markup.warningN s);
+    Output.Private_Hooks.error_fn :=
+      (fn (i, s) => standard_message mbox (SOME i) Isabelle_Markup.errorN s);
+    Output.Private_Hooks.protocol_message_fn := message true mbox Isabelle_Markup.protocolN;
     Output.Private_Hooks.urgent_message_fn := ! Output.Private_Hooks.writeln_fn;
     Output.Private_Hooks.prompt_fn := ignore;
-    message true mbox "A" [] (Session.welcome ())
+    message true mbox Isabelle_Markup.initN [] (Session.welcome ())
   end;
 
 end;