--- a/src/HOL/ATP_Linkup.thy Thu Dec 20 13:58:45 2007 +0100
+++ b/src/HOL/ATP_Linkup.thy Thu Dec 20 14:33:40 2007 +0100
@@ -8,7 +8,7 @@
theory ATP_Linkup
imports PreList Hilbert_Choice
- (*It must be a parent or a child of every other theory, to prevent theory-merge errors.*)
+ (*FIXME It must be a parent or a child of every other theory, to prevent theory-merge errors. FIXME*)
uses
"Tools/polyhash.ML"
"Tools/res_clause.ML"
@@ -19,6 +19,7 @@
("Tools/res_atp.ML")
("Tools/res_atp_provers.ML")
("Tools/res_atp_methods.ML")
+ "~~/src/Tools/random_word.ML"
"~~/src/Tools/Metis/metis.ML"
("Tools/metis_tools.ML")
begin
--- a/src/HOL/IsaMakefile Thu Dec 20 13:58:45 2007 +0100
+++ b/src/HOL/IsaMakefile Thu Dec 20 14:33:40 2007 +0100
@@ -87,11 +87,11 @@
$(SRC)/Provers/order.ML \
$(SRC)/Provers/project_rule.ML $(SRC)/Provers/quantifier1.ML \
$(SRC)/Provers/quasi.ML $(SRC)/Provers/splitter.ML \
- $(SRC)/Provers/trancl.ML $(SRC)/Tools/Metis/metis.ML\
- $(SRC)/Tools/code/code_funcgr.ML \
- $(SRC)/Tools/code/code_name.ML $(SRC)/Tools/code/code_package.ML \
- $(SRC)/Tools/code/code_target.ML $(SRC)/Tools/code/code_thingol.ML \
- $(SRC)/Tools/nbe.ML $(SRC)/Tools/rat.ML Tools/TFL/casesplit.ML ATP_Linkup.thy \
+ $(SRC)/Provers/trancl.ML $(SRC)/Tools/Metis/metis.ML \
+ $(SRC)/Tools/code/code_funcgr.ML $(SRC)/Tools/code/code_name.ML \
+ $(SRC)/Tools/code/code_package.ML $(SRC)/Tools/code/code_target.ML \
+ $(SRC)/Tools/code/code_thingol.ML $(SRC)/Tools/nbe.ML \
+ $(SRC)/Tools/random_word.ML $(SRC)/Tools/rat.ML Tools/TFL/casesplit.ML ATP_Linkup.thy \
Accessible_Part.thy Arith_Tools.thy Code_Setup.thy Datatype.thy \
Dense_Linear_Order.thy Divides.thy Equiv_Relations.thy Extraction.thy \
Finite_Set.thy Fun.thy FunDef.thy HOL.thy \
--- a/src/Pure/General/ROOT.ML Thu Dec 20 13:58:45 2007 +0100
+++ b/src/Pure/General/ROOT.ML Thu Dec 20 14:33:40 2007 +0100
@@ -16,7 +16,6 @@
use "../ML/ml_parse.ML";
use "secure.ML";
-use "random_word.ML";
use "integer.ML";
use "stack.ML";
use "heap.ML";
--- a/src/Pure/General/random_word.ML Thu Dec 20 13:58:45 2007 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,39 +0,0 @@
-(* Title: Pure/General/random_word.ML
- ID: $Id$
- Author: Makarius
-
-Simple generator for pseudo-random numbers, using unboxed word
-arithmetic only. Unprotected concurrency introduces some true
-randomness.
-*)
-
-signature RANDOM_WORD =
-sig
- val range: int
- val range_real: real
- val next: unit -> word
- val next_bit: unit -> bool
-end;
-
-structure RandomWord: RANDOM_WORD =
-struct
-
-(*minimum length of unboxed words on all supported ML platforms*)
-val _ = Word.wordSize >= 30 orelse raise Fail ("Bad platform word size");
-
-val range = 0x40000000;
-val range_real = 1073741824.0;
-val mask = Word.fromInt (range - 1);
-val max_bit = Word.fromInt (range div 2);
-
-(*multiplier according to Borosh and Niederreiter (for m = 2^30),
- cf. http://random.mat.sbg.ac.at/~charly/server/server.html*)
-val a = 0w777138309;
-
-(*generator*)
-val rand = ref 0w0;
-fun next () = (rand := Word.andb (a * ! rand + 0w1, mask); ! rand);
-fun next_bit () = Word.andb (next (), max_bit) = 0w0;
-
-end;
-
--- a/src/Pure/IsaMakefile Thu Dec 20 13:58:45 2007 +0100
+++ b/src/Pure/IsaMakefile Thu Dec 20 14:33:40 2007 +0100
@@ -29,12 +29,12 @@
General/heap.ML General/history.ML General/integer.ML General/markup.ML \
General/name_space.ML General/ord_list.ML General/output.ML \
General/path.ML General/position.ML General/pretty.ML \
- General/print_mode.ML General/random_word.ML General/scan.ML \
- General/secure.ML General/seq.ML General/source.ML General/stack.ML \
- General/susp.ML General/symbol.ML General/table.ML General/url.ML \
- General/xml.ML Isar/ROOT.ML Isar/antiquote.ML Isar/args.ML Isar/attrib.ML \
- Isar/auto_bind.ML Isar/calculation.ML Isar/class.ML Isar/code.ML \
- Isar/code_unit.ML Isar/constdefs.ML Isar/context_rules.ML Isar/element.ML \
+ General/print_mode.ML General/scan.ML General/secure.ML General/seq.ML \
+ General/source.ML General/stack.ML General/susp.ML General/symbol.ML \
+ General/table.ML General/url.ML General/xml.ML Isar/ROOT.ML \
+ Isar/antiquote.ML Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML \
+ Isar/calculation.ML Isar/class.ML Isar/code.ML Isar/code_unit.ML \
+ Isar/constdefs.ML Isar/context_rules.ML Isar/element.ML \
Isar/find_theorems.ML Isar/instance.ML Isar/isar_cmd.ML \
Isar/isar_syn.ML Isar/local_defs.ML Isar/local_syntax.ML \
Isar/local_theory.ML Isar/locale.ML Isar/method.ML Isar/net_rules.ML \
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/random_word.ML Thu Dec 20 14:33:40 2007 +0100
@@ -0,0 +1,39 @@
+(* Title: Tools/random_word.ML
+ ID: $Id$
+ Author: Makarius
+
+Simple generator for pseudo-random numbers, using unboxed word
+arithmetic only. Unprotected concurrency introduces some true
+randomness.
+*)
+
+signature RANDOM_WORD =
+sig
+ val range: int
+ val range_real: real
+ val next: unit -> word
+ val next_bit: unit -> bool
+end;
+
+structure RandomWord: RANDOM_WORD =
+struct
+
+(*minimum length of unboxed words on all supported ML platforms*)
+val _ = Word.wordSize >= 30 orelse raise Fail ("Bad platform word size");
+
+val range = 0x40000000;
+val range_real = 1073741824.0;
+val mask = Word.fromInt (range - 1);
+val max_bit = Word.fromInt (range div 2);
+
+(*multiplier according to Borosh and Niederreiter (for m = 2^30),
+ cf. http://random.mat.sbg.ac.at/~charly/server/server.html*)
+val a = 0w777138309;
+
+(*generator*)
+val rand = ref 0w0;
+fun next () = (rand := Word.andb (a * ! rand + 0w1, mask); ! rand);
+fun next_bit () = Word.andb (next (), max_bit) = 0w0;
+
+end;
+