mesh facts by taking into consideration whether a fact is known to MeSh
authorblanchet
Wed, 18 Jul 2012 08:44:04 +0200
changeset 48313 0faafdffa662
parent 48312 b40722a81ac9
child 48314 ee33ba3c0e05
mesh facts by taking into consideration whether a fact is known to MeSh
src/HOL/TPTP/mash_eval.ML
src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
--- a/src/HOL/TPTP/mash_eval.ML	Wed Jul 18 08:44:04 2012 +0200
+++ b/src/HOL/TPTP/mash_eval.ML	Wed Jul 18 08:44:04 2012 +0200
@@ -30,6 +30,7 @@
     val {provers, max_facts, slice, type_enc, lam_trans, timeout, ...} =
       Sledgehammer_Isar.default_params ctxt []
     val prover_name = hd provers
+    val slack_max_facts = max_facts_slack * the max_facts
     val path = file_name |> Path.explode
     val lines = path |> File.read_lines
     val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
@@ -39,10 +40,6 @@
     val mash_ok = Unsynchronized.ref 0
     val mesh_ok = Unsynchronized.ref 0
     val isar_ok = Unsynchronized.ref 0
-    fun sugg_facts hyp_ts concl_t suggs =
-      suggested_facts suggs
-      #> chop (max_facts_slack * the max_facts)
-      #>> Sledgehammer_Fact.maybe_instantiate_inducts ctxt hyp_ts concl_t
     fun with_index facts s = (find_index (curry (op =) s) facts + 1, s)
     fun index_string (j, s) = s ^ "@" ^ string_of_int j
     fun str_of_res label facts {outcome, run_time, used_facts, ...} =
@@ -64,12 +61,6 @@
                   |> map index_string
                   |> space_implode " "))
       end
-    fun prove ok heading facts goal =
-      let
-        val facts = facts |> take (the max_facts)
-        val res as {outcome, ...} = run_prover ctxt params facts goal
-        val _ = if is_none outcome then ok := !ok + 1 else ()
-      in str_of_res heading facts res end
     fun solve_goal (j, line) =
       let
         val (name, suggs) = extract_query line
@@ -77,23 +68,29 @@
           case find_first (fn (_, th) => Thm.get_name_hint th = name) facts of
             SOME (_, th) => th
           | NONE => error ("No fact called \"" ^ name ^ "\"")
-        val isar_deps = isabelle_dependencies_of all_names th
         val goal = goal_of_thm thy th
         val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
+        val isar_deps = isabelle_dependencies_of all_names th
         val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
-        val (isar_facts, _) = facts |> sugg_facts hyp_ts concl_t isar_deps
+        val isar_facts = suggested_facts isar_deps facts
         val iter_facts =
           Sledgehammer_Filter_Iter.iterative_relevant_facts ctxt params
-              prover_name (max_facts_slack * the max_facts) NONE hyp_ts concl_t
-              facts
-        val (mash_facts, mash_rejects) =
-          facts |> sugg_facts hyp_ts concl_t suggs
-        val mesh_facts =
-          mesh_facts (the max_facts) iter_facts mash_facts mash_rejects
-        val iter_s = prove iter_ok iterN iter_facts goal
-        val mash_s = prove mash_ok mashN mash_facts goal
-        val mesh_s = prove mesh_ok meshN mesh_facts goal
-        val isar_s = prove isar_ok isarN isar_facts goal
+              prover_name slack_max_facts NONE hyp_ts concl_t facts
+        val mash_facts = suggested_facts suggs facts
+        val mess = [(iter_facts, SOME (length iter_facts)), (mash_facts, NONE)]
+        val mesh_facts = mesh_facts slack_max_facts mess
+        fun prove ok heading facts =
+          let
+            val facts =
+              facts |> Sledgehammer_Fact.maybe_instantiate_inducts ctxt hyp_ts concl_t
+                    |> take (the max_facts)
+            val res as {outcome, ...} = run_prover ctxt params facts goal
+            val _ = if is_none outcome then ok := !ok + 1 else ()
+          in str_of_res heading facts res end
+        val iter_s = prove iter_ok iterN iter_facts
+        val mash_s = prove mash_ok mashN mash_facts
+        val mesh_s = prove mesh_ok meshN mesh_facts
+        val isar_s = prove isar_ok isarN isar_facts
       in
         ["Goal " ^ string_of_int j ^ ": " ^ name, iter_s, mash_s, mesh_s,
          isar_s]
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML	Wed Jul 18 08:44:04 2012 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML	Wed Jul 18 08:44:04 2012 +0200
@@ -21,7 +21,7 @@
   val unescape_metas : string -> string list
   val extract_query : string -> string * string list
   val suggested_facts : string list -> fact list -> fact list
-  val mesh_facts : int -> fact list -> fact list -> fact list -> fact list
+  val mesh_facts : int -> (fact list * int option) list -> fact list
   val all_non_tautological_facts_of :
     theory -> status Termtab.table -> fact list
   val theory_ord : theory * theory -> order
@@ -106,21 +106,22 @@
   find_first (fn (_, th) => Thm.get_name_hint th = sugg) facts
 fun suggested_facts suggs facts = map_filter (find_suggested facts) suggs
 
-fun mesh_facts max_facts iter_facts mash_facts mash_rejects =
+fun sum_avg n xs =
+  fold (Integer.add o Integer.mult n) xs 0 div (length xs)
+
+fun mesh_facts max_facts mess =
   let
+    val n = length mess
     val fact_eq = Thm.eq_thm o pairself snd
-    val num_iter = length iter_facts
-    val num_mash = length mash_facts
-    fun score_in f fs n =
-      case find_index (curry fact_eq f) fs of
-        ~1 => length fs
-      | j => j
-    fun score_of fact =
-      score_in fact iter_facts num_iter + score_in fact mash_facts num_mash
+    fun score_in fact (facts, def) =
+      case find_index (curry fact_eq fact) facts of
+        ~1 => def
+      | j => SOME j
+    fun score_of fact = mess |> map_filter (score_in fact) |> sum_avg n
+    val facts = fold (union fact_eq o take max_facts o fst) mess []
   in
-    union fact_eq iter_facts mash_facts
-    |> map (`score_of) |> sort (int_ord o pairself fst) |> map snd
-    |> take max_facts
+    facts |> map (`score_of) |> sort (int_ord o pairself fst) |> map snd
+          |> take max_facts
   end
 
 val thy_feature_prefix = "y_"
@@ -320,8 +321,8 @@
     val command =
       getenv "MASH_HOME" ^ "/mash.py --inputFile " ^ cmd_file ^
       " --outputDir " ^ mash_dir () ^ " --predictions " ^ pred_file ^
-      " --log " ^ log_file ^ (if save then " --saveModel" else "") ^
-      " > /dev/null"
+      " --log " ^ log_file ^ " --numberOfPredictions 1000" ^
+      (if save then " --saveModel" else "") ^ " > /dev/null"
     val _ = File.write cmd_path ""
     val _ = write_cmds (File.append cmd_path)
     val _ = trace_msg ctxt (fn () => "  running " ^ command)
@@ -524,11 +525,11 @@
       val iter_facts =
         iterative_relevant_facts ctxt params prover max_facts NONE hyp_ts
                                  concl_t facts
-      val (mash_facts, mash_rejects) =
+      val mash_facts =
         facts |> mash_suggest_facts ctxt params prover hyp_ts concl_t
-              |> chop max_facts
+      val mess = [(iter_facts, SOME (length iter_facts)), (mash_facts, NONE)]
     in
-      mesh_facts max_facts iter_facts mash_facts mash_rejects
+      mesh_facts max_facts mess
       |> not (null add_ths) ? prepend_facts add_ths
     end