src/HOL/TPTP/mash_eval.ML
author blanchet
Fri, 20 Jul 2012 22:19:46 +0200
changeset 49421 b002cc16aa99
parent 49419 0a261b4aa093
child 49453 3e45c98fe127
permissions -rw-r--r--
honor suggested MaSh weights
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@49330
    19
open Sledgehammer_Fact
blanchet@49396
    20
open Sledgehammer_MaSh
blanchet@49255
    21
blanchet@49394
    22
val MePoN = "MePo"
blanchet@49394
    23
val MaShN = "MaSh"
blanchet@49394
    24
val MeshN = "Mesh"
blanchet@49394
    25
val IsarN = "Isar"
blanchet@49256
    26
blanchet@49308
    27
val max_facts_slack = 2
blanchet@49250
    28
blanchet@49339
    29
val all_names =
blanchet@49348
    30
  filter_out is_likely_tautology_or_too_meta
blanchet@49393
    31
  #> map (rpair () o nickname_of) #> Symtab.make
blanchet@49331
    32
blanchet@49300
    33
fun evaluate_mash_suggestions ctxt params thy file_name =
blanchet@49250
    34
  let
blanchet@49308
    35
    val {provers, max_facts, slice, type_enc, lam_trans, timeout, ...} =
blanchet@49254
    36
      Sledgehammer_Isar.default_params ctxt []
blanchet@49333
    37
    val prover = hd provers
blanchet@49328
    38
    val slack_max_facts = max_facts_slack * the max_facts
blanchet@49250
    39
    val path = file_name |> Path.explode
blanchet@49250
    40
    val lines = path |> File.read_lines
blanchet@49314
    41
    val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
blanchet@49409
    42
    val facts = all_facts_of (Proof_Context.init_global thy) css_table
blanchet@49339
    43
    val all_names = all_names (facts |> map snd)
blanchet@49394
    44
    val mepo_ok = Unsynchronized.ref 0
blanchet@49256
    45
    val mash_ok = Unsynchronized.ref 0
blanchet@49313
    46
    val mesh_ok = Unsynchronized.ref 0
blanchet@49315
    47
    val isar_ok = Unsynchronized.ref 0
blanchet@49304
    48
    fun with_index facts s = (find_index (curry (op =) s) facts + 1, s)
blanchet@49250
    49
    fun index_string (j, s) = s ^ "@" ^ string_of_int j
blanchet@49255
    50
    fun str_of_res label facts {outcome, run_time, used_facts, ...} =
blanchet@49304
    51
      let val facts = facts |> map (fn ((name, _), _) => name ()) in
blanchet@49304
    52
        "  " ^ label ^ ": " ^
blanchet@49304
    53
        (if is_none outcome then
blanchet@49304
    54
           "Success (" ^ ATP_Util.string_from_time run_time ^ "): " ^
blanchet@49304
    55
           (used_facts |> map (with_index facts o fst)
blanchet@49304
    56
                       |> sort (int_ord o pairself fst)
blanchet@49304
    57
                       |> map index_string
blanchet@49304
    58
                       |> space_implode " ") ^
blanchet@49308
    59
           (if length facts < the max_facts then
blanchet@49304
    60
              " (of " ^ string_of_int (length facts) ^ ")"
blanchet@49304
    61
            else
blanchet@49304
    62
              "")
blanchet@49304
    63
         else
blanchet@49304
    64
           "Failure: " ^
blanchet@49308
    65
           (facts |> take (the max_facts) |> tag_list 1
blanchet@49304
    66
                  |> map index_string
blanchet@49304
    67
                  |> space_implode " "))
blanchet@49304
    68
      end
blanchet@49326
    69
    fun solve_goal (j, line) =
blanchet@49250
    70
      let
blanchet@49326
    71
        val (name, suggs) = extract_query line
blanchet@49250
    72
        val th =
blanchet@49393
    73
          case find_first (fn (_, th) => nickname_of th = name) facts of
blanchet@49250
    74
            SOME (_, th) => th
blanchet@49250
    75
          | NONE => error ("No fact called \"" ^ name ^ "\"")
blanchet@49265
    76
        val goal = goal_of_thm thy th
blanchet@49260
    77
        val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
blanchet@49419
    78
        val isar_deps = isar_dependencies_of all_names th |> these
blanchet@49250
    79
        val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
blanchet@49394
    80
        val mepo_facts =
blanchet@49421
    81
          Sledgehammer_MePo.mepo_suggested_facts ctxt params prover
blanchet@49396
    82
              slack_max_facts NONE hyp_ts concl_t facts
blanchet@49421
    83
          |> Sledgehammer_MePo.weight_mepo_facts
blanchet@49421
    84
        val mash_facts = suggested_facts suggs facts
blanchet@49394
    85
        val mess = [(mepo_facts, []), (mash_facts, [])]
blanchet@49328
    86
        val mesh_facts = mesh_facts slack_max_facts mess
blanchet@49421
    87
        val isar_facts = suggested_facts (map (rpair 1.0) isar_deps) facts
blanchet@49421
    88
        fun prove ok heading get facts =
blanchet@49328
    89
          let
blanchet@49328
    90
            val facts =
blanchet@49421
    91
              facts |> map get
blanchet@49421
    92
                    |> Sledgehammer_Fact.maybe_instantiate_inducts ctxt hyp_ts
blanchet@49421
    93
                                                                   concl_t
blanchet@49421
    94
                    |> take (the max_facts)
blanchet@49336
    95
            val res as {outcome, ...} =
blanchet@49336
    96
              run_prover_for_mash ctxt params prover facts goal
blanchet@49328
    97
            val _ = if is_none outcome then ok := !ok + 1 else ()
blanchet@49328
    98
          in str_of_res heading facts res end
blanchet@49421
    99
        val mepo_s = prove mepo_ok MePoN fst mepo_facts
blanchet@49421
   100
        val mash_s = prove mash_ok MaShN fst mash_facts
blanchet@49421
   101
        val mesh_s = prove mesh_ok MeshN I mesh_facts
blanchet@49421
   102
        val isar_s = prove isar_ok IsarN fst isar_facts
blanchet@49250
   103
      in
blanchet@49394
   104
        ["Goal " ^ string_of_int j ^ ": " ^ name, mepo_s, mash_s, mesh_s,
blanchet@49315
   105
         isar_s]
blanchet@49256
   106
        |> cat_lines |> tracing
blanchet@49250
   107
      end
blanchet@49256
   108
    fun total_of heading ok n =
blanchet@49394
   109
      "  " ^ heading ^ ": " ^ string_of_int (!ok) ^ " (" ^
blanchet@49256
   110
      Real.fmt (StringCvt.FIX (SOME 1))
blanchet@49256
   111
               (100.0 * Real.fromInt (!ok) / Real.fromInt n) ^ "%)"
blanchet@49265
   112
    val inst_inducts = Config.get ctxt Sledgehammer_Fact.instantiate_inducts
blanchet@49260
   113
    val options =
blanchet@49333
   114
      [prover, string_of_int (the max_facts) ^ " facts",
blanchet@49256
   115
       "slice" |> not slice ? prefix "dont_", the_default "smart" type_enc,
blanchet@49256
   116
       the_default "smart" lam_trans, ATP_Util.string_from_time timeout,
blanchet@49256
   117
       "instantiate_inducts" |> not inst_inducts ? prefix "dont_"]
blanchet@49256
   118
    val n = length lines
blanchet@49256
   119
  in
blanchet@49256
   120
    tracing " * * *";
blanchet@49260
   121
    tracing ("Options: " ^ commas options);
blanchet@49326
   122
    List.app solve_goal (tag_list 1 lines);
blanchet@49256
   123
    ["Successes (of " ^ string_of_int n ^ " goals)",
blanchet@49394
   124
     total_of MePoN mepo_ok n,
blanchet@49394
   125
     total_of MaShN mash_ok n,
blanchet@49394
   126
     total_of MeshN mesh_ok n,
blanchet@49394
   127
     total_of IsarN isar_ok n]
blanchet@49256
   128
    |> cat_lines |> tracing
blanchet@49256
   129
  end
blanchet@49250
   130
blanchet@49249
   131
end;