src/HOL/Tools/Sledgehammer/sledgehammer_util.ML
changeset 36142 f5e15e9aae10
parent 36062 194cb6e3c13f
child 36169 27b1cc58715e
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Wed Apr 14 17:10:16 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Wed Apr 14 18:23:51 2010 +0200
@@ -6,6 +6,7 @@
 
 signature SLEDGEHAMMER_UTIL =
 sig
+  val plural_s : int -> string
   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
@@ -17,12 +18,7 @@
 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 plural_s n = if n = 1 then "" else "s"
 
 fun serial_commas _ [] = ["??"]
   | serial_commas _ [s] = [s]
@@ -60,4 +56,11 @@
         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;