src/HOL/TPTP/mash_eval.ML
author blanchet
Wed, 18 Jul 2012 08:44:04 +0200
changeset 49326 3c4e10606567
parent 49322 7c78f14d20cf
child 49327 b40722a81ac9
permissions -rw-r--r--
implemented MaSh QUERY operation
blanchet@49300
     1
(*  Title:      HOL/TPTP/mash_eval.ML
blanchet@49249
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@49249
     3
    Copyright   2012
blanchet@49249
     4
blanchet@49300
     5
Evaluate proof suggestions from MaSh (Machine-learning for Sledgehammer).
blanchet@49249
     6
*)
blanchet@49249
     7
blanchet@49300
     8
signature MASH_EVAL =
blanchet@49249
     9
sig
blanchet@49266
    10
  type params = Sledgehammer_Provers.params
blanchet@49266
    11
blanchet@49300
    12
  val evaluate_mash_suggestions :
blanchet@49266
    13
    Proof.context -> params -> theory -> string -> unit
blanchet@49249
    14
end;
blanchet@49249
    15
blanchet@49300
    16
structure MaSh_Eval : MASH_EVAL =
blanchet@49249
    17
struct
blanchet@49250
    18
blanchet@49266
    19
open Sledgehammer_Filter_MaSh
blanchet@49255
    20
blanchet@49315
    21
val isarN = "Isa"
blanchet@49300
    22
val iterN = "Iter"
blanchet@49256
    23
val mashN = "MaSh"
blanchet@49313
    24
val meshN = "Mesh"
blanchet@49256
    25
blanchet@49308
    26
val max_facts_slack = 2
blanchet@49250
    27
blanchet@49300
    28
fun evaluate_mash_suggestions ctxt params thy file_name =
blanchet@49250
    29
  let
blanchet@49308
    30
    val {provers, max_facts, slice, type_enc, lam_trans, timeout, ...} =
blanchet@49254
    31
      Sledgehammer_Isar.default_params ctxt []
blanchet@49254
    32
    val prover_name = hd provers
blanchet@49250
    33
    val path = file_name |> Path.explode
blanchet@49250
    34
    val lines = path |> File.read_lines
blanchet@49314
    35
    val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
blanchet@49314
    36
    val facts = all_non_tautological_facts_of thy css_table
blanchet@49250
    37
    val all_names = facts |> map (Thm.get_name_hint o snd)
blanchet@49266
    38
    val iter_ok = Unsynchronized.ref 0
blanchet@49256
    39
    val mash_ok = Unsynchronized.ref 0
blanchet@49313
    40
    val mesh_ok = Unsynchronized.ref 0
blanchet@49315
    41
    val isar_ok = Unsynchronized.ref 0
blanchet@49326
    42
    fun sugg_facts hyp_ts concl_t suggs =
blanchet@49326
    43
      suggested_facts suggs
blanchet@49308
    44
      #> take (max_facts_slack * the max_facts)
blanchet@49265
    45
      #> Sledgehammer_Fact.maybe_instantiate_inducts ctxt hyp_ts concl_t
blanchet@49313
    46
    fun mesh_facts fsp =
blanchet@49250
    47
      let
blanchet@49304
    48
        val (fs1, fs2) =
blanchet@49304
    49
          fsp |> pairself (map (apfst (apfst (fn name => name ()))))
blanchet@49250
    50
        val fact_eq = (op =) o pairself fst
blanchet@49250
    51
        fun score_in f fs =
blanchet@49250
    52
          case find_index (curry fact_eq f) fs of
blanchet@49250
    53
            ~1 => length fs
blanchet@49250
    54
          | j => j
blanchet@49250
    55
        fun score_of f = score_in f fs1 + score_in f fs2
blanchet@49250
    56
      in
blanchet@49250
    57
        union fact_eq fs1 fs2
blanchet@49251
    58
        |> map (`score_of) |> sort (int_ord o pairself fst) |> map snd
blanchet@49308
    59
        |> take (the max_facts)
blanchet@49304
    60
        |> map (apfst (apfst K))
blanchet@49250
    61
      end
blanchet@49304
    62
    fun with_index facts s = (find_index (curry (op =) s) facts + 1, s)
blanchet@49250
    63
    fun index_string (j, s) = s ^ "@" ^ string_of_int j
blanchet@49255
    64
    fun str_of_res label facts {outcome, run_time, used_facts, ...} =
blanchet@49304
    65
      let val facts = facts |> map (fn ((name, _), _) => name ()) in
blanchet@49304
    66
        "  " ^ label ^ ": " ^
blanchet@49304
    67
        (if is_none outcome then
blanchet@49304
    68
           "Success (" ^ ATP_Util.string_from_time run_time ^ "): " ^
blanchet@49304
    69
           (used_facts |> map (with_index facts o fst)
blanchet@49304
    70
                       |> sort (int_ord o pairself fst)
blanchet@49304
    71
                       |> map index_string
blanchet@49304
    72
                       |> space_implode " ") ^
blanchet@49308
    73
           (if length facts < the max_facts then
blanchet@49304
    74
              " (of " ^ string_of_int (length facts) ^ ")"
blanchet@49304
    75
            else
blanchet@49304
    76
              "")
blanchet@49304
    77
         else
blanchet@49304
    78
           "Failure: " ^
blanchet@49308
    79
           (facts |> take (the max_facts) |> tag_list 1
blanchet@49304
    80
                  |> map index_string
blanchet@49304
    81
                  |> space_implode " "))
blanchet@49304
    82
      end
blanchet@49266
    83
    fun prove ok heading facts goal =
blanchet@49254
    84
      let
blanchet@49308
    85
        val facts = facts |> take (the max_facts)
blanchet@49266
    86
        val res as {outcome, ...} = run_prover ctxt params facts goal
blanchet@49256
    87
        val _ = if is_none outcome then ok := !ok + 1 else ()
blanchet@49256
    88
      in str_of_res heading facts res end
blanchet@49326
    89
    fun solve_goal (j, line) =
blanchet@49250
    90
      let
blanchet@49326
    91
        val (name, suggs) = extract_query line
blanchet@49250
    92
        val th =
blanchet@49250
    93
          case find_first (fn (_, th) => Thm.get_name_hint th = name) facts of
blanchet@49250
    94
            SOME (_, th) => th
blanchet@49250
    95
          | NONE => error ("No fact called \"" ^ name ^ "\"")
blanchet@49315
    96
        val isar_deps = isabelle_dependencies_of all_names th
blanchet@49265
    97
        val goal = goal_of_thm thy th
blanchet@49260
    98
        val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
blanchet@49250
    99
        val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
blanchet@49326
   100
        val isar_facts = facts |> sugg_facts hyp_ts concl_t isar_deps
blanchet@49266
   101
        val iter_facts =
blanchet@49307
   102
          Sledgehammer_Filter_Iter.iterative_relevant_facts ctxt params
blanchet@49308
   103
              prover_name (max_facts_slack * the max_facts) NONE hyp_ts concl_t
blanchet@49308
   104
              facts
blanchet@49326
   105
        val mash_facts = facts |> sugg_facts hyp_ts concl_t suggs
blanchet@49313
   106
        val mesh_facts = mesh_facts (iter_facts, mash_facts)
blanchet@49266
   107
        val iter_s = prove iter_ok iterN iter_facts goal
blanchet@49266
   108
        val mash_s = prove mash_ok mashN mash_facts goal
blanchet@49313
   109
        val mesh_s = prove mesh_ok meshN mesh_facts goal
blanchet@49315
   110
        val isar_s = prove isar_ok isarN isar_facts goal
blanchet@49250
   111
      in
blanchet@49315
   112
        ["Goal " ^ string_of_int j ^ ": " ^ name, iter_s, mash_s, mesh_s,
blanchet@49315
   113
         isar_s]
blanchet@49256
   114
        |> cat_lines |> tracing
blanchet@49250
   115
      end
blanchet@49256
   116
    fun total_of heading ok n =
blanchet@49256
   117
      " " ^ heading ^ ": " ^ string_of_int (!ok) ^ " (" ^
blanchet@49256
   118
      Real.fmt (StringCvt.FIX (SOME 1))
blanchet@49256
   119
               (100.0 * Real.fromInt (!ok) / Real.fromInt n) ^ "%)"
blanchet@49265
   120
    val inst_inducts = Config.get ctxt Sledgehammer_Fact.instantiate_inducts
blanchet@49260
   121
    val options =
blanchet@49308
   122
      [prover_name, string_of_int (the max_facts) ^ " facts",
blanchet@49256
   123
       "slice" |> not slice ? prefix "dont_", the_default "smart" type_enc,
blanchet@49256
   124
       the_default "smart" lam_trans, ATP_Util.string_from_time timeout,
blanchet@49256
   125
       "instantiate_inducts" |> not inst_inducts ? prefix "dont_"]
blanchet@49256
   126
    val n = length lines
blanchet@49256
   127
  in
blanchet@49256
   128
    tracing " * * *";
blanchet@49260
   129
    tracing ("Options: " ^ commas options);
blanchet@49326
   130
    List.app solve_goal (tag_list 1 lines);
blanchet@49256
   131
    ["Successes (of " ^ string_of_int n ^ " goals)",
blanchet@49266
   132
     total_of iterN iter_ok n,
blanchet@49256
   133
     total_of mashN mash_ok n,
blanchet@49313
   134
     total_of meshN mesh_ok n,
blanchet@49315
   135
     total_of isarN isar_ok n]
blanchet@49256
   136
    |> cat_lines |> tracing
blanchet@49256
   137
  end
blanchet@49250
   138
blanchet@49249
   139
end;