src/HOL/Tools/Sledgehammer/sledgehammer_util.ML
changeset 36378 f32c567dbcaa
parent 36183 f723348b2231
child 36402 1b20805974c7
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Fri Apr 23 18:11:41 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Fri Apr 23 19:12:49 2010 +0200
@@ -6,6 +6,7 @@
 
 signature SLEDGEHAMMER_UTIL =
 sig
+  val pairf : ('a -> 'b) -> ('a -> 'c) -> 'a -> 'b * 'c
   val plural_s : int -> string
   val serial_commas : string -> string list -> string list
   val replace_all : string -> string -> string -> string
@@ -13,14 +14,13 @@
   val timestamp : unit -> string
   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
 
+fun pairf f g x = (f x, g x)
+
 fun plural_s n = if n = 1 then "" else "s"
 
 fun serial_commas _ [] = ["??"]
@@ -73,11 +73,4 @@
         SOME (Time.fromMilliseconds msecs)
     end
 
-(* 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
-
 end;