src/HOL/Tools/Sledgehammer/sledgehammer_util.ML
changeset 36062 194cb6e3c13f
parent 35963 943e2582dc87
child 36142 f5e15e9aae10
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Mon Mar 29 15:26:19 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Mon Mar 29 15:50:18 2010 +0200
@@ -1,4 +1,4 @@
-(*  Title:      HOL/Sledgehammer/sledgehammer_util.ML
+(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_util.ML
     Author:     Jasmin Blanchette, TU Muenchen
 
 General-purpose functions used by the Sledgehammer modules.
@@ -9,11 +9,21 @@
   val serial_commas : string -> string list -> string list
   val parse_bool_option : bool -> string -> string -> bool option
   val parse_time_option : string -> string -> Time.time option
+  val hashw : word * word -> word
+  val hashw_char : Char.char * word -> word
+  val hashw_string : string * word -> word
 end;
 
 structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
 struct
 
+(* This hash function is recommended in Compilers: Principles, Techniques, and
+   Tools, by Aho, Sethi and Ullman. The hashpjw function, which they
+   particularly recommend, triggers a bug in versions of Poly/ML up to 4.2.0. *)
+fun hashw (u, w) = Word.+ (u, Word.* (0w65599, w))
+fun hashw_char (c, w) = hashw (Word.fromInt (Char.ord c), w)
+fun hashw_string (s:string, w) = CharVector.foldl hashw_char w s
+
 fun serial_commas _ [] = ["??"]
   | serial_commas _ [s] = [s]
   | serial_commas conj [s1, s2] = [s1, conj, s2]