src/HOL/Tools/Sledgehammer/sledgehammer_util.ML
author blanchet
Wed, 11 Jul 2012 21:43:19 +0200
changeset 48251 6cdcfbddc077
parent 46957 0c15caf47040
child 48292 7fcee834c7f5
permissions -rw-r--r--
moved most of MaSh exporter code to Sledgehammer
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36062
194cb6e3c13f get rid of Polyhash, since it's no longer used
blanchet
parents: 35963
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_util.ML
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
     3
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
     4
General-purpose functions used by the Sledgehammer modules.
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
     5
*)
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
     6
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
     7
signature SLEDGEHAMMER_UTIL =
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
     8
sig
36142
f5e15e9aae10 make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents: 36062
diff changeset
     9
  val plural_s : int -> string
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    10
  val serial_commas : string -> string list -> string list
38738
0ce517c1970f make sure that "undo_ascii_of" is the inverse of "ascii_of", also for non-printable characters -- and avoid those in ``-style facts
blanchet
parents: 38698
diff changeset
    11
  val simplify_spaces : string -> string
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    12
  val parse_bool_option : bool -> string -> string -> bool option
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    13
  val parse_time_option : string -> string -> Time.time option
38044
463177795c49 minor refactoring
blanchet
parents: 38019
diff changeset
    14
  val subgoal_count : Proof.state -> int
43043
1406f6fc5dc3 normalize indices in chained facts to make sure that backtick facts (which often result in different names) are recognized + changed definition of urgent messages
blanchet
parents: 43034
diff changeset
    15
  val normalize_chained_theorems : thm list -> thm list
38696
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38652
diff changeset
    16
  val reserved_isar_keyword_table : unit -> unit Symtab.table
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    17
  val thms_in_proof : string list option -> thm -> string list
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    18
end;
39318
ad9a1f9b0558 implemented Auto Sledgehammer
blanchet
parents: 38818
diff changeset
    19
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    20
structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    21
struct
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    22
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43043
diff changeset
    23
open ATP_Util
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43043
diff changeset
    24
36142
f5e15e9aae10 make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents: 36062
diff changeset
    25
fun plural_s n = if n = 1 then "" else "s"
36062
194cb6e3c13f get rid of Polyhash, since it's no longer used
blanchet
parents: 35963
diff changeset
    26
43029
3e060b1c844b use helpers and tweak Quickcheck's priority to it comes second (to give Solve Direct slightly more time before another prover runs)
blanchet
parents: 43005
diff changeset
    27
val serial_commas = Try.serial_commas
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43043
diff changeset
    28
val simplify_spaces = strip_spaces false (K true)
37962
d7dbe01f48d7 keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents: 37575
diff changeset
    29
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    30
fun parse_bool_option option name s =
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    31
  (case s of
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    32
     "smart" => if option then NONE else raise Option
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    33
   | "false" => SOME false
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    34
   | "true" => SOME true
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    35
   | "" => SOME true
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    36
   | _ => raise Option)
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    37
  handle Option.Option =>
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    38
         let val ss = map quote ((option ? cons "smart") ["true", "false"]) in
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    39
           error ("Parameter " ^ quote name ^ " must be assigned " ^
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    40
                  space_implode " " (serial_commas "or" ss) ^ ".")
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    41
         end
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    42
40341
03156257040f standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents: 39890
diff changeset
    43
val has_junk =
40627
becf5d5187cc renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents: 40341
diff changeset
    44
  exists (fn s => not (Symbol.is_digit s) andalso s <> ".") o raw_explode (* FIXME Symbol.explode (?) *)
40341
03156257040f standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents: 39890
diff changeset
    45
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    46
fun parse_time_option _ "none" = NONE
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    47
  | parse_time_option name s =
40341
03156257040f standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents: 39890
diff changeset
    48
    let val secs = if has_junk s then NONE else Real.fromString s in
43034
18259246abb5 try both "metis" and (on failure) "metisFT" in replay
blanchet
parents: 43033
diff changeset
    49
      if is_none secs orelse Real.< (the secs, 0.0) then
18259246abb5 try both "metis" and (on failure) "metisFT" in replay
blanchet
parents: 43033
diff changeset
    50
        error ("Parameter " ^ quote name ^ " must be assigned a nonnegative \
40341
03156257040f standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents: 39890
diff changeset
    51
               \number of seconds (e.g., \"60\", \"0.5\") or \"none\".")
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    52
      else
40341
03156257040f standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents: 39890
diff changeset
    53
        SOME (seconds (the secs))
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    54
    end
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    55
43029
3e060b1c844b use helpers and tweak Quickcheck's priority to it comes second (to give Solve Direct slightly more time before another prover runs)
blanchet
parents: 43005
diff changeset
    56
val subgoal_count = Try.subgoal_count
38044
463177795c49 minor refactoring
blanchet
parents: 38019
diff changeset
    57
43043
1406f6fc5dc3 normalize indices in chained facts to make sure that backtick facts (which often result in different names) are recognized + changed definition of urgent messages
blanchet
parents: 43034
diff changeset
    58
val normalize_chained_theorems =
1406f6fc5dc3 normalize indices in chained facts to make sure that backtick facts (which often result in different names) are recognized + changed definition of urgent messages
blanchet
parents: 43034
diff changeset
    59
  maps (fn th => insert Thm.eq_thm_prop (zero_var_indexes th) [th])
38696
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38652
diff changeset
    60
fun reserved_isar_keyword_table () =
46957
0c15caf47040 clarified Keyword.is_keyword: union of minor and major;
wenzelm
parents: 43085
diff changeset
    61
  Keyword.dest () |-> union (op =)
38696
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38652
diff changeset
    62
  |> map (rpair ()) |> Symtab.make
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38652
diff changeset
    63
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    64
(* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    65
   fixes that seem to be missing over there; or maybe the two code portions are
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    66
   not doing the same? *)
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    67
fun fold_body_thms thm_name f =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    68
  let
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    69
    fun app n (PBody {thms, ...}) =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    70
      thms |> fold (fn (_, (name, prop, body)) => fn x =>
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    71
        let
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    72
          val body' = Future.join body
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    73
          val n' =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    74
            n + (if name = "" orelse
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    75
                    (* uncommon case where the proved theorem occurs twice
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    76
                       (e.g., "Transitive_Closure.trancl_into_trancl") *)
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    77
                    (n = 1 andalso name = thm_name) then
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    78
                   0
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    79
                 else
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    80
                   1)
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    81
          val x' = x |> n' <= 1 ? app n' body'
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    82
        in (x' |> n = 1 ? f (name, prop, body')) end)
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    83
  in fold (app 0) end
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    84
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    85
fun thms_in_proof all_names th =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    86
  let
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    87
    val is_name_ok =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    88
      case all_names of
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    89
        SOME names => member (op =) names
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    90
      | NONE => (fn s => s <> "" andalso not (String.isPrefix "Pure." s))
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    91
    fun collect (s, _, _) = is_name_ok s ? insert (op =) s
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    92
    val names =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    93
      [] |> fold_body_thms (Thm.get_name_hint th) collect [Thm.proof_body_of th]
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    94
  in names end
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 46957
diff changeset
    95
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents:
diff changeset
    96
end;