src/HOL/Tools/Sledgehammer/sledgehammer_util.ML
author blanchet
Fri, 16 Apr 2010 15:49:13 +0200
changeset 36170 0cdb76723c88
parent 36169 27b1cc58715e
child 36183 f723348b2231
permissions -rw-r--r--
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_util.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3 
     4 General-purpose functions used by the Sledgehammer modules.
     5 *)
     6 
     7 signature SLEDGEHAMMER_UTIL =
     8 sig
     9   val plural_s : int -> string
    10   val serial_commas : string -> string list -> string list
    11   val replace_all : string -> string -> string -> string
    12   val remove_all : string -> string -> string
    13   val timestamp : unit -> string
    14   val parse_bool_option : bool -> string -> string -> bool option
    15   val parse_time_option : string -> string -> Time.time option
    16   val hashw : word * word -> word
    17   val hashw_char : Char.char * word -> word
    18   val hashw_string : string * word -> word
    19 end;
    20  
    21 structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
    22 struct
    23 
    24 fun plural_s n = if n = 1 then "" else "s"
    25 
    26 fun serial_commas _ [] = ["??"]
    27   | serial_commas _ [s] = [s]
    28   | serial_commas conj [s1, s2] = [s1, conj, s2]
    29   | serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3]
    30   | serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss
    31 
    32 fun replace_all bef aft =
    33   let
    34     fun aux seen "" = String.implode (rev seen)
    35       | aux seen s =
    36       if String.isPrefix bef s then
    37           aux seen "" ^ aft ^ aux [] (unprefix bef s)
    38         else
    39           aux (String.sub (s, 0) :: seen) (String.extract (s, 1, NONE))
    40   in aux [] end
    41 
    42 fun remove_all bef = replace_all bef ""
    43 
    44 val timestamp = Date.fmt "%Y-%m-%d %H:%M:%S" o Date.fromTimeLocal o Time.now
    45 
    46 fun parse_bool_option option name s =
    47   (case s of
    48      "smart" => if option then NONE else raise Option
    49    | "false" => SOME false
    50    | "true" => SOME true
    51    | "" => SOME true
    52    | _ => raise Option)
    53   handle Option.Option =>
    54          let val ss = map quote ((option ? cons "smart") ["true", "false"]) in
    55            error ("Parameter " ^ quote name ^ " must be assigned " ^
    56                   space_implode " " (serial_commas "or" ss) ^ ".")
    57          end
    58 
    59 fun parse_time_option _ "none" = NONE
    60   | parse_time_option name s =
    61     let
    62       val msecs =
    63         case space_explode " " s of
    64           [s1, "min"] => 60000 * the (Int.fromString s1)
    65         | [s1, "s"] => 1000 * the (Int.fromString s1)
    66         | [s1, "ms"] => the (Int.fromString s1)
    67         | _ => 0
    68     in
    69       if msecs <= 0 then
    70         error ("Parameter " ^ quote name ^ " must be assigned a positive time \
    71                \value (e.g., \"60 s\", \"200 ms\") or \"none\".")
    72       else
    73         SOME (Time.fromMilliseconds msecs)
    74     end
    75 
    76 (* This hash function is recommended in Compilers: Principles, Techniques, and
    77    Tools, by Aho, Sethi and Ullman. The hashpjw function, which they
    78    particularly recommend, triggers a bug in versions of Poly/ML up to 4.2.0. *)
    79 fun hashw (u, w) = Word.+ (u, Word.* (0w65599, w))
    80 fun hashw_char (c, w) = hashw (Word.fromInt (Char.ord c), w)
    81 fun hashw_string (s:string, w) = CharVector.foldl hashw_char w s
    82 
    83 end;