src/HOL/TPTP/mash_import.ML
author blanchet
Wed, 11 Jul 2012 21:43:19 +0200
changeset 49265 1065c307fafe
parent 49260 854a47677335
child 49266 6cdcfbddc077
permissions -rw-r--r--
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet@49249
     1
(*  Title:      HOL/TPTP/mash_import.ML
blanchet@49249
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@49249
     3
    Copyright   2012
blanchet@49249
     4
blanchet@49249
     5
Import proof suggestions from MaSh (Machine-learning for Sledgehammer) and
blanchet@49249
     6
evaluate them.
blanchet@49249
     7
*)
blanchet@49249
     8
blanchet@49249
     9
signature MASH_IMPORT =
blanchet@49249
    10
sig
blanchet@49250
    11
  val import_and_evaluate_mash_suggestions :
blanchet@49250
    12
    Proof.context -> theory -> string -> unit
blanchet@49249
    13
end;
blanchet@49249
    14
blanchet@49249
    15
structure MaSh_Import : MASH_IMPORT =
blanchet@49249
    16
struct
blanchet@49250
    17
blanchet@49255
    18
open MaSh_Export
blanchet@49255
    19
blanchet@49250
    20
val unescape_meta =
blanchet@49250
    21
  let
blanchet@49250
    22
    fun un accum [] = String.implode (rev accum)
blanchet@49250
    23
      | un accum (#"\\" :: d1 :: d2 :: d3 :: cs) =
blanchet@49250
    24
        (case Int.fromString (String.implode [d1, d2, d3]) of
blanchet@49250
    25
           SOME n => un (Char.chr n :: accum) cs
blanchet@49250
    26
         | NONE => "" (* error *))
blanchet@49250
    27
      | un _ (#"\\" :: _) = "" (* error *)
blanchet@49250
    28
      | un accum (c :: cs) = un (c :: accum) cs
blanchet@49250
    29
  in un [] o String.explode end
blanchet@49250
    30
blanchet@49250
    31
val of_fact_name = unescape_meta
blanchet@49250
    32
blanchet@49256
    33
val depN = "Dependencies"
blanchet@49256
    34
val mengN = "Meng & Paulson"
blanchet@49256
    35
val mashN = "MaSh"
blanchet@49256
    36
val meng_mashN = "M&P+MaSh"
blanchet@49256
    37
blanchet@49254
    38
val max_relevant_slack = 2
blanchet@49250
    39
blanchet@49250
    40
fun import_and_evaluate_mash_suggestions ctxt thy file_name =
blanchet@49250
    41
  let
blanchet@49260
    42
    val {provers, max_relevant, slice, type_enc, lam_trans, timeout, ...} =
blanchet@49254
    43
      Sledgehammer_Isar.default_params ctxt []
blanchet@49254
    44
    val prover_name = hd provers
blanchet@49250
    45
    val path = file_name |> Path.explode
blanchet@49250
    46
    val lines = path |> File.read_lines
blanchet@49250
    47
    val facts = non_tautological_facts_of thy
blanchet@49250
    48
    val all_names = facts |> map (Thm.get_name_hint o snd)
blanchet@49256
    49
    val meng_ok = Unsynchronized.ref 0
blanchet@49256
    50
    val mash_ok = Unsynchronized.ref 0
blanchet@49256
    51
    val meng_mash_ok = Unsynchronized.ref 0
blanchet@49256
    52
    val dep_ok = Unsynchronized.ref 0
blanchet@49250
    53
    fun find_sugg facts sugg =
blanchet@49250
    54
      find_first (fn (_, th) => Thm.get_name_hint th = sugg) facts
blanchet@49251
    55
    fun sugg_facts hyp_ts concl_t facts =
blanchet@49250
    56
      map_filter (find_sugg facts o of_fact_name)
blanchet@49254
    57
      #> take (max_relevant_slack * the max_relevant)
blanchet@49265
    58
      #> Sledgehammer_Fact.maybe_instantiate_inducts ctxt hyp_ts concl_t
blanchet@49250
    59
      #> map (apfst (apfst (fn name => name ())))
blanchet@49256
    60
    fun meng_mash_facts fs1 fs2 =
blanchet@49250
    61
      let
blanchet@49250
    62
        val fact_eq = (op =) o pairself fst
blanchet@49250
    63
        fun score_in f fs =
blanchet@49250
    64
          case find_index (curry fact_eq f) fs of
blanchet@49250
    65
            ~1 => length fs
blanchet@49250
    66
          | j => j
blanchet@49250
    67
        fun score_of f = score_in f fs1 + score_in f fs2
blanchet@49250
    68
      in
blanchet@49250
    69
        union fact_eq fs1 fs2
blanchet@49251
    70
        |> map (`score_of) |> sort (int_ord o pairself fst) |> map snd
blanchet@49254
    71
        |> take (the max_relevant)
blanchet@49250
    72
      end
blanchet@49250
    73
    fun with_index facts s =
blanchet@49250
    74
      (find_index (fn ((s', _), _) => s = s') facts + 1, s)
blanchet@49250
    75
    fun index_string (j, s) = s ^ "@" ^ string_of_int j
blanchet@49255
    76
    fun str_of_res label facts {outcome, run_time, used_facts, ...} =
blanchet@49255
    77
      "  " ^ label ^ ": " ^
blanchet@49255
    78
      (if is_none outcome then
blanchet@49255
    79
         "Success (" ^ ATP_Util.string_from_time run_time ^ "): " ^
blanchet@49256
    80
         (used_facts |> map (with_index facts o fst)
blanchet@49256
    81
                     |> sort (int_ord o pairself fst)
blanchet@49256
    82
                     |> map index_string
blanchet@49256
    83
                     |> space_implode " ") ^
blanchet@49256
    84
         (if length facts < the max_relevant then
blanchet@49256
    85
            " (of " ^ string_of_int (length facts) ^ ")"
blanchet@49256
    86
          else
blanchet@49256
    87
            "")
blanchet@49255
    88
       else
blanchet@49256
    89
         "Failure: " ^
blanchet@49256
    90
         (facts |> map (fst o fst)
blanchet@49260
    91
                |> take (the max_relevant)
blanchet@49256
    92
                |> tag_list 1
blanchet@49256
    93
                |> map index_string
blanchet@49256
    94
                |> space_implode " "))
blanchet@49256
    95
    fun run_sh ok heading facts goal =
blanchet@49254
    96
      let
blanchet@49260
    97
        val facts = facts |> take (the max_relevant)
blanchet@49260
    98
        val res as {outcome, ...} = run_sledgehammer ctxt facts goal
blanchet@49256
    99
        val _ = if is_none outcome then ok := !ok + 1 else ()
blanchet@49256
   100
      in str_of_res heading facts res end
blanchet@49256
   101
    fun solve_goal j name suggs =
blanchet@49250
   102
      let
blanchet@49250
   103
        val name = of_fact_name name
blanchet@49250
   104
        val th =
blanchet@49250
   105
          case find_first (fn (_, th) => Thm.get_name_hint th = name) facts of
blanchet@49250
   106
            SOME (_, th) => th
blanchet@49250
   107
          | NONE => error ("No fact called \"" ^ name ^ "\"")
blanchet@49250
   108
        val deps = dependencies_of all_names th
blanchet@49265
   109
        val goal = goal_of_thm thy th
blanchet@49260
   110
        val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
blanchet@49250
   111
        val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
blanchet@49251
   112
        val deps_facts = sugg_facts hyp_ts concl_t facts deps
blanchet@49254
   113
        val meng_facts =
blanchet@49260
   114
          meng_paulson_facts ctxt (max_relevant_slack * the max_relevant) goal
blanchet@49260
   115
                             facts
blanchet@49251
   116
        val mash_facts = sugg_facts hyp_ts concl_t facts suggs
blanchet@49256
   117
        val meng_mash_facts = meng_mash_facts meng_facts mash_facts
blanchet@49256
   118
        val meng_s = run_sh meng_ok mengN meng_facts goal
blanchet@49256
   119
        val mash_s = run_sh mash_ok mashN mash_facts goal
blanchet@49256
   120
        val meng_mash_s = run_sh meng_mash_ok meng_mashN meng_mash_facts goal
blanchet@49256
   121
        val dep_s = run_sh dep_ok depN deps_facts goal
blanchet@49250
   122
      in
blanchet@49256
   123
        ["Goal " ^ string_of_int j ^ ": " ^ name, meng_s, mash_s, meng_mash_s,
blanchet@49256
   124
         dep_s]
blanchet@49256
   125
        |> cat_lines |> tracing
blanchet@49250
   126
      end
blanchet@49250
   127
    val explode_suggs = space_explode " " #> filter_out (curry (op =) "")
blanchet@49256
   128
    fun do_line (j, line) =
blanchet@49250
   129
      case space_explode ":" line of
blanchet@49256
   130
        [goal_name, suggs] => solve_goal j goal_name (explode_suggs suggs)
blanchet@49250
   131
      | _ => ()
blanchet@49256
   132
    fun total_of heading ok n =
blanchet@49256
   133
      " " ^ heading ^ ": " ^ string_of_int (!ok) ^ " (" ^
blanchet@49256
   134
      Real.fmt (StringCvt.FIX (SOME 1))
blanchet@49256
   135
               (100.0 * Real.fromInt (!ok) / Real.fromInt n) ^ "%)"
blanchet@49265
   136
    val inst_inducts = Config.get ctxt Sledgehammer_Fact.instantiate_inducts
blanchet@49260
   137
    val options =
blanchet@49256
   138
      [prover_name, string_of_int (the max_relevant) ^ " facts",
blanchet@49256
   139
       "slice" |> not slice ? prefix "dont_", the_default "smart" type_enc,
blanchet@49256
   140
       the_default "smart" lam_trans, ATP_Util.string_from_time timeout,
blanchet@49256
   141
       "instantiate_inducts" |> not inst_inducts ? prefix "dont_"]
blanchet@49256
   142
    val n = length lines
blanchet@49256
   143
  in
blanchet@49256
   144
    tracing " * * *";
blanchet@49260
   145
    tracing ("Options: " ^ commas options);
blanchet@49256
   146
    List.app do_line (tag_list 1 lines);
blanchet@49256
   147
    ["Successes (of " ^ string_of_int n ^ " goals)",
blanchet@49256
   148
     total_of mengN meng_ok n,
blanchet@49256
   149
     total_of mashN mash_ok n,
blanchet@49256
   150
     total_of meng_mashN meng_mash_ok n,
blanchet@49256
   151
     total_of depN dep_ok n]
blanchet@49256
   152
    |> cat_lines |> tracing
blanchet@49256
   153
  end
blanchet@49250
   154
blanchet@49249
   155
end;