--- a/src/HOL/Mirabelle/Mirabelle.thy Mon Sep 06 12:38:45 2010 +0200
+++ b/src/HOL/Mirabelle/Mirabelle.thy Mon Sep 06 13:06:27 2010 +0200
@@ -7,7 +7,7 @@
uses "Tools/mirabelle.ML"
begin
-(* no multithreading, no parallel proofs *)
+(* no multithreading, no parallel proofs *) (* FIXME *)
ML {* Multithreading.max_threads := 1 *}
ML {* Goal.parallel_proofs := 0 *}
--- a/src/HOL/Tools/Nitpick/nitpick_isar.ML Mon Sep 06 12:38:45 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_isar.ML Mon Sep 06 13:06:27 2010 +0200
@@ -348,7 +348,7 @@
|>> curry (op =) "genuine")
in
if auto orelse blocking then go ()
- else (Toplevel.thread true (tap go); (false, state))
+ else (Toplevel.thread true (tap go); (false, state)) (* FIXME no threads in user-space *)
end
fun nitpick_trans (params, i) =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML Mon Sep 06 12:38:45 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML Mon Sep 06 13:06:27 2010 +0200
@@ -452,6 +452,7 @@
else
()
end
+ (* FIXME no threads in user-space *)
in if blocking then run () else Toplevel.thread true (tap run) |> K () end
val setup =
--- a/src/HOL/ex/ROOT.ML Mon Sep 06 12:38:45 2010 +0200
+++ b/src/HOL/ex/ROOT.ML Mon Sep 06 13:06:27 2010 +0200
@@ -70,7 +70,7 @@
HTML.with_charset "utf-8" (no_document use_thys)
["Hebrew", "Chinese", "Serbian"];
-(setmp_noncritical proofs 2 (setmp_noncritical Multithreading.max_threads 1 use_thy))
+(setmp_noncritical proofs 2 (setmp_noncritical Multithreading.max_threads 1 use_thy)) (* FIXME *)
"Hilbert_Classical";
use_thy "SVC_Oracle";
--- a/src/Tools/WWW_Find/scgi_server.ML Mon Sep 06 12:38:45 2010 +0200
+++ b/src/Tools/WWW_Find/scgi_server.ML Mon Sep 06 13:06:27 2010 +0200
@@ -50,7 +50,7 @@
else (tracing ("Waiting for a free thread...");
ConditionVar.wait (thread_wait, thread_wait_mutex));
add_thread
- (Thread.fork
+ (Thread.fork (* FIXME avoid low-level Poly/ML thread primitives *)
(fn () => exception_trace threadf,
[Thread.EnableBroadcastInterrupt true,
Thread.InterruptState