src/HOL/Tools/Sledgehammer/sledgehammer_util.ML
author blanchet
Fri, 20 Jul 2012 22:19:46 +0200
changeset 49407 ca998fa08cd9
parent 49398 df75b2d7e26a
child 49671 5caa414ce9a2
permissions -rw-r--r--
added "learn_from_atp" command to MaSh, for patient users
     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 sledgehammerN : string
    10   val plural_s : int -> string
    11   val serial_commas : string -> string list -> string list
    12   val simplify_spaces : string -> string
    13   val infinite_timeout : Time.time
    14   val time_mult : real -> Time.time -> Time.time
    15   val parse_bool_option : bool -> string -> string -> bool option
    16   val parse_time_option : string -> string -> Time.time option
    17   val subgoal_count : Proof.state -> int
    18   val reserved_isar_keyword_table : unit -> unit Symtab.table
    19   val thms_in_proof : unit Symtab.table option -> thm -> string list
    20 end;
    21 
    22 structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
    23 struct
    24 
    25 open ATP_Util
    26 
    27 val sledgehammerN = "sledgehammer"
    28 
    29 fun plural_s n = if n = 1 then "" else "s"
    30 
    31 val serial_commas = Try.serial_commas
    32 val simplify_spaces = strip_spaces false (K true)
    33 
    34 val infinite_timeout = seconds 31536000.0 (* one year *)
    35 
    36 fun time_mult k t =
    37   Time.fromMilliseconds (Real.ceil (k * Real.fromInt (Time.toMilliseconds t)))
    38 
    39 fun parse_bool_option option name s =
    40   (case s of
    41      "smart" => if option then NONE else raise Option
    42    | "false" => SOME false
    43    | "true" => SOME true
    44    | "" => SOME true
    45    | _ => raise Option)
    46   handle Option.Option =>
    47          let val ss = map quote ((option ? cons "smart") ["true", "false"]) in
    48            error ("Parameter " ^ quote name ^ " must be assigned " ^
    49                   space_implode " " (serial_commas "or" ss) ^ ".")
    50          end
    51 
    52 val has_junk =
    53   exists (fn s => not (Symbol.is_digit s) andalso s <> ".") o raw_explode (* FIXME Symbol.explode (?) *)
    54 
    55 fun parse_time_option _ "none" = NONE
    56   | parse_time_option name s =
    57     let val secs = if has_junk s then NONE else Real.fromString s in
    58       if is_none secs orelse Real.< (the secs, 0.0) then
    59         error ("Parameter " ^ quote name ^ " must be assigned a nonnegative \
    60                \number of seconds (e.g., \"60\", \"0.5\") or \"none\".")
    61       else
    62         SOME (seconds (the secs))
    63     end
    64 
    65 val subgoal_count = Try.subgoal_count
    66 
    67 fun reserved_isar_keyword_table () =
    68   Keyword.dest () |-> union (op =) |> map (rpair ()) |> Symtab.make
    69 
    70 (* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few
    71    fixes that seem to be missing over there; or maybe the two code portions are
    72    not doing the same? *)
    73 fun fold_body_thms thm_name f =
    74   let
    75     fun app n (PBody {thms, ...}) =
    76       thms |> fold (fn (_, (name, _, body)) => fn accum =>
    77           let
    78             (* The second disjunct caters for the uncommon case where the proved
    79                theorem occurs in its own proof (e.g.,
    80                "Transitive_Closure.trancl_into_trancl"). *)
    81             val no_name = (name = "" orelse (n = 1 andalso name = thm_name))
    82             val accum =
    83               accum |> (if n = 1 andalso not no_name then f name else I)
    84             val n = n + (if no_name then 0 else 1)
    85           in accum |> (if n <= 1 then app n (Future.join body) else I) end)
    86   in fold (app 0) end
    87 
    88 fun thms_in_proof all_names th =
    89   let
    90     val collect =
    91       case all_names of
    92         SOME names => (fn s => Symtab.defined names s ? insert (op =) s)
    93       | NONE => insert (op =)
    94     val names =
    95       [] |> fold_body_thms (Thm.get_name_hint th) collect [Thm.proof_body_of th]
    96   in names end
    97 
    98 end;