| author | wenzelm | 
| Sat, 27 Jul 2013 16:44:58 +0200 | |
| changeset 52734 | 077149654ab4 | 
| parent 52196 | 2281f33e8da6 | 
| child 53139 | 07a6e11f1631 | 
| permissions | -rw-r--r-- | 
| 48285 | 1 | (* Title: HOL/TPTP/mash_eval.ML | 
| 48234 | 2 | Author: Jasmin Blanchette, TU Muenchen | 
| 3 | Copyright 2012 | |
| 4 | ||
| 48285 | 5 | Evaluate proof suggestions from MaSh (Machine-learning for Sledgehammer). | 
| 48234 | 6 | *) | 
| 7 | ||
| 48285 | 8 | signature MASH_EVAL = | 
| 48234 | 9 | sig | 
| 48251 
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
 blanchet parents: 
48250diff
changeset | 10 | type params = Sledgehammer_Provers.params | 
| 
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
 blanchet parents: 
48250diff
changeset | 11 | |
| 51015 | 12 | val MePoN : string | 
| 50953 | 13 | val MaSh_IsarN : string | 
| 14 | val MaSh_ProverN : string | |
| 15 | val MeSh_IsarN : string | |
| 16 | val MeSh_ProverN : string | |
| 17 | val IsarN : string | |
| 50437 | 18 | val evaluate_mash_suggestions : | 
| 51182 
962190eab40d
provide two modes for MaSh driver: linearized or real visibility
 blanchet parents: 
51135diff
changeset | 19 | Proof.context -> params -> int * int option -> bool -> string list | 
| 
962190eab40d
provide two modes for MaSh driver: linearized or real visibility
 blanchet parents: 
51135diff
changeset | 20 | -> string option -> string -> string -> string -> string -> string -> string | 
| 
962190eab40d
provide two modes for MaSh driver: linearized or real visibility
 blanchet parents: 
51135diff
changeset | 21 | -> unit | 
| 48234 | 22 | end; | 
| 23 | ||
| 48285 | 24 | structure MaSh_Eval : MASH_EVAL = | 
| 48234 | 25 | struct | 
| 48235 | 26 | |
| 50557 | 27 | open Sledgehammer_Util | 
| 48315 
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
 blanchet parents: 
48313diff
changeset | 28 | open Sledgehammer_Fact | 
| 50557 | 29 | open Sledgehammer_MePo | 
| 48381 | 30 | open Sledgehammer_MaSh | 
| 50557 | 31 | open Sledgehammer_Provers | 
| 32 | open Sledgehammer_Isar | |
| 48240 | 33 | |
| 51008 | 34 | val MaSh_IsarN = MaShN ^ "-Isar" | 
| 35 | val MaSh_ProverN = MaShN ^ "-Prover" | |
| 36 | val MeSh_IsarN = MeShN ^ "-Isar" | |
| 37 | val MeSh_ProverN = MeShN ^ "-Prover" | |
| 48379 
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
 blanchet parents: 
48378diff
changeset | 38 | val IsarN = "Isar" | 
| 48241 | 39 | |
| 50559 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 40 | fun in_range (from, to) j = | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 41 | j >= from andalso (to = NONE orelse j <= the to) | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 42 | |
| 51182 
962190eab40d
provide two modes for MaSh driver: linearized or real visibility
 blanchet parents: 
51135diff
changeset | 43 | fun evaluate_mash_suggestions ctxt params range linearize methods prob_dir_name | 
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 44 | mepo_file_name mash_isar_file_name mash_prover_file_name | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 45 | mesh_isar_file_name mesh_prover_file_name report_file_name = | 
| 48235 | 46 | let | 
| 51028 | 47 | val zeros = [0, 0, 0, 0, 0, 0] | 
| 50448 
0a740d127187
have MaSh evaluator keep all raw problem/solution files in a directory
 blanchet parents: 
50442diff
changeset | 48 | val report_path = report_file_name |> Path.explode | 
| 
0a740d127187
have MaSh evaluator keep all raw problem/solution files in a directory
 blanchet parents: 
50442diff
changeset | 49 | val _ = File.write report_path "" | 
| 50589 | 50 | fun print s = File.append report_path (s ^ "\n") | 
| 48293 | 51 |     val {provers, max_facts, slice, type_enc, lam_trans, timeout, ...} =
 | 
| 50557 | 52 | default_params ctxt [] | 
| 48318 | 53 | val prover = hd provers | 
| 50412 | 54 | val slack_max_facts = generous_max_facts (the max_facts) | 
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 55 | val lines_of = Path.explode #> try File.read_lines #> these | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 56 | val file_names = | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 57 | [mepo_file_name, mash_isar_file_name, mash_prover_file_name, | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 58 | mesh_isar_file_name, mesh_prover_file_name] | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 59 | val lines as [mepo_lines, mash_isar_lines, mash_prover_lines, | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 60 | mesh_isar_lines, mesh_prover_lines] = | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 61 | map lines_of file_names | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 62 | val num_lines = fold (Integer.max o length) lines 0 | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 63 | fun pad lines = lines @ replicate (num_lines - length lines) "" | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 64 | val lines = | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 65 | pad mepo_lines ~~ pad mash_isar_lines ~~ pad mash_prover_lines ~~ | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 66 | pad mesh_isar_lines ~~ pad mesh_prover_lines | 
| 50557 | 67 | val css = clasimpset_rule_table_of ctxt | 
| 50442 
4f6a4d32522c
don't blacklist "case" theorems -- this causes problems in MaSh later
 blanchet parents: 
50440diff
changeset | 68 | val facts = all_facts ctxt true false Symtab.empty [] [] css | 
| 50735 
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
 blanchet parents: 
50633diff
changeset | 69 | val name_tabs = build_name_tables nickname_of_thm facts | 
| 48289 | 70 | fun with_index facts s = (find_index (curry (op =) s) facts + 1, s) | 
| 50952 | 71 | fun index_str (j, s) = s ^ "@" ^ string_of_int j | 
| 50953 | 72 | val str_of_method = enclose " " ": " | 
| 73 |     fun str_of_result method facts ({outcome, run_time, used_facts, ...}
 | |
| 50952 | 74 | : prover_result) = | 
| 51004 
5f2788c38127
distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
 blanchet parents: 
50967diff
changeset | 75 | let val facts = facts |> map (fst o fst) in | 
| 50953 | 76 | str_of_method method ^ | 
| 48289 | 77 | (if is_none outcome then | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51199diff
changeset | 78 |            "Success (" ^ ATP_Util.string_of_time run_time ^ "): " ^
 | 
| 48289 | 79 | (used_facts |> map (with_index facts o fst) | 
| 80 | |> sort (int_ord o pairself fst) | |
| 50952 | 81 | |> map index_str | 
| 48289 | 82 | |> space_implode " ") ^ | 
| 48293 | 83 | (if length facts < the max_facts then | 
| 48289 | 84 | " (of " ^ string_of_int (length facts) ^ ")" | 
| 85 | else | |
| 86 | "") | |
| 87 | else | |
| 88 | "Failure: " ^ | |
| 48293 | 89 | (facts |> take (the max_facts) |> tag_list 1 | 
| 50952 | 90 | |> map index_str | 
| 48289 | 91 | |> space_implode " ")) | 
| 92 | end | |
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 93 | fun solve_goal (j, ((((mepo_line, mash_isar_line), mash_prover_line), | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 94 | mesh_isar_line), mesh_prover_line)) = | 
| 50559 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 95 | if in_range range j then | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 96 | let | 
| 50965 | 97 | val get_suggs = extract_suggestions ##> take slack_max_facts | 
| 98 | val (name1, mepo_suggs) = get_suggs mepo_line | |
| 99 | val (name2, mash_isar_suggs) = get_suggs mash_isar_line | |
| 100 | val (name3, mash_prover_suggs) = get_suggs mash_prover_line | |
| 101 | val (name4, mesh_isar_suggs) = get_suggs mesh_isar_line | |
| 102 | val (name5, mesh_prover_suggs) = get_suggs mesh_prover_line | |
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 103 | val [name] = | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 104 | [name1, name2, name3, name4, name5] | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 105 | |> filter (curry (op <>) "") |> distinct (op =) | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 106 | handle General.Match => error "Input files out of sync." | 
| 50559 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 107 | val th = | 
| 50624 
4d0997abce79
improved thm order hack, in case the default names are overridden
 blanchet parents: 
50620diff
changeset | 108 | case find_first (fn (_, th) => nickname_of_thm th = name) facts of | 
| 50559 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 109 | SOME (_, th) => th | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 110 |             | NONE => error ("No fact called \"" ^ name ^ "\".")
 | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 111 | val goal = goal_of_thm (Proof_Context.theory_of ctxt) th | 
| 52196 
2281f33e8da6
redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
 blanchet parents: 
52125diff
changeset | 112 | val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal 1 ctxt | 
| 50754 
74a6adcb96ac
also generate queries for goals with too many Isar dependencies
 blanchet parents: 
50735diff
changeset | 113 | val isar_deps = isar_dependencies_of name_tabs th | 
| 51135 | 114 | val facts = | 
| 51199 
e3447c380fe1
honor linearization option also in the evaluation driver
 blanchet parents: 
51182diff
changeset | 115 | facts | 
| 
e3447c380fe1
honor linearization option also in the evaluation driver
 blanchet parents: 
51182diff
changeset | 116 | |> filter (fn (_, th') => | 
| 
e3447c380fe1
honor linearization option also in the evaluation driver
 blanchet parents: 
51182diff
changeset | 117 | if linearize then crude_thm_ord (th', th) = LESS | 
| 
e3447c380fe1
honor linearization option also in the evaluation driver
 blanchet parents: 
51182diff
changeset | 118 | else thm_less (th', th)) | 
| 51134 | 119 | val find_suggs = | 
| 120 | find_suggested_facts ctxt facts #> map fact_of_raw_fact | |
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 121 | fun get_facts [] compute = compute facts | 
| 50967 
00d87ade2294
optimization -- evaluate conversion to table only once
 blanchet parents: 
50965diff
changeset | 122 | | get_facts suggs _ = find_suggs suggs | 
| 50559 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 123 | val mepo_facts = | 
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 124 | get_facts mepo_suggs (fn _ => | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 125 | mepo_suggested_facts ctxt params prover slack_max_facts NONE | 
| 50965 | 126 | hyp_ts concl_t facts) | 
| 127 | |> weight_mepo_facts | |
| 50952 | 128 | fun mash_of suggs = | 
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 129 | get_facts suggs (fn _ => | 
| 51134 | 130 | find_mash_suggestions ctxt slack_max_facts suggs facts [] [] | 
| 51004 
5f2788c38127
distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
 blanchet parents: 
50967diff
changeset | 131 | |> fst |> map fact_of_raw_fact) | 
| 50965 | 132 | |> weight_mash_facts | 
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 133 | val mash_isar_facts = mash_of mash_isar_suggs | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 134 | val mash_prover_facts = mash_of mash_prover_suggs | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 135 | fun mess_of mash_facts = | 
| 50814 | 136 | [(mepo_weight, (mepo_facts, [])), | 
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 137 | (mash_weight, (mash_facts, []))] | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 138 | fun mesh_of suggs mash_facts = | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 139 | get_facts suggs (fn _ => | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 140 | mesh_facts (Thm.eq_thm_prop o pairself snd) slack_max_facts | 
| 50965 | 141 | (mess_of mash_facts)) | 
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 142 | val mesh_isar_facts = mesh_of mesh_isar_suggs mash_isar_facts | 
| 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 143 | val mesh_prover_facts = mesh_of mesh_prover_suggs mash_prover_facts | 
| 50967 
00d87ade2294
optimization -- evaluate conversion to table only once
 blanchet parents: 
50965diff
changeset | 144 | val isar_facts = find_suggs isar_deps | 
| 50559 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 145 | (* adapted from "mirabelle_sledgehammer.ML" *) | 
| 50953 | 146 | fun set_file_name method (SOME dir) = | 
| 50559 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 147 | let | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 148 | val prob_prefix = | 
| 50826 | 149 | "goal_" ^ string_of_int j ^ "__" ^ encode_str name ^ "__" ^ | 
| 50953 | 150 | method | 
| 50559 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 151 | in | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 152 | Config.put dest_dir dir | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 153 | #> Config.put problem_prefix (prob_prefix ^ "__") | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 154 | #> Config.put SMT_Config.debug_files (dir ^ "/" ^ prob_prefix) | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 155 | end | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 156 | | set_file_name _ NONE = I | 
| 50965 | 157 | fun prove method get facts = | 
| 50953 | 158 | if not (member (op =) methods method) orelse | 
| 159 | (null facts andalso method <> IsarN) then | |
| 160 | (str_of_method method ^ "Skipped", 0) | |
| 50952 | 161 | else | 
| 162 | let | |
| 163 | fun nickify ((_, stature), th) = | |
| 164 | ((K (encode_str (nickname_of_thm th)), stature), th) | |
| 165 | val facts = | |
| 166 | facts | |
| 50965 | 167 | |> map (get #> nickify) | 
| 50952 | 168 | |> maybe_instantiate_inducts ctxt hyp_ts concl_t | 
| 169 | |> take (the max_facts) | |
| 51004 
5f2788c38127
distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
 blanchet parents: 
50967diff
changeset | 170 | |> map fact_of_raw_fact | 
| 50953 | 171 | val ctxt = ctxt |> set_file_name method prob_dir_name | 
| 50952 | 172 |                 val res as {outcome, ...} =
 | 
| 173 | run_prover_for_mash ctxt params prover facts goal | |
| 174 | val ok = if is_none outcome then 1 else 0 | |
| 50953 | 175 | in (str_of_result method facts res, ok) end | 
| 50591 | 176 | val ress = | 
| 50965 | 177 | [fn () => prove MePoN fst mepo_facts, | 
| 178 | fn () => prove MaSh_IsarN fst mash_isar_facts, | |
| 179 | fn () => prove MaSh_ProverN fst mash_prover_facts, | |
| 180 | fn () => prove MeSh_IsarN I mesh_isar_facts, | |
| 181 | fn () => prove MeSh_ProverN I mesh_prover_facts, | |
| 182 | fn () => prove IsarN I isar_facts] | |
| 50559 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 183 | |> (* Par_List. *) map (fn f => f ()) | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 184 | in | 
| 50591 | 185 | "Goal " ^ string_of_int j ^ ": " ^ name :: map fst ress | 
| 50587 
bd6582be1562
synchronize access to shared reference and print proper total
 blanchet parents: 
50563diff
changeset | 186 | |> cat_lines |> print; | 
| 50591 | 187 | map snd ress | 
| 50559 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 188 | end | 
| 
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
 blanchet parents: 
50557diff
changeset | 189 | else | 
| 51028 | 190 | zeros | 
| 50953 | 191 | fun total_of method ok n = | 
| 192 |       str_of_method method ^ string_of_int ok ^ " (" ^
 | |
| 50591 | 193 | Real.fmt (StringCvt.FIX (SOME 1)) | 
| 51028 | 194 | (100.0 * Real.fromInt ok / Real.fromInt (Int.max (1, n))) ^ "%)" | 
| 50557 | 195 | val inst_inducts = Config.get ctxt instantiate_inducts | 
| 48245 | 196 | val options = | 
| 50767 | 197 | ["prover = " ^ prover, | 
| 198 | "max_facts = " ^ string_of_int (the max_facts), | |
| 199 | "slice" |> not slice ? prefix "dont_", | |
| 200 | "type_enc = " ^ the_default "smart" type_enc, | |
| 201 | "lam_trans = " ^ the_default "smart" lam_trans, | |
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51199diff
changeset | 202 | "timeout = " ^ ATP_Util.string_of_time (the_default one_year timeout), | 
| 48241 | 203 | "instantiate_inducts" |> not inst_inducts ? prefix "dont_"] | 
| 50587 
bd6582be1562
synchronize access to shared reference and print proper total
 blanchet parents: 
50563diff
changeset | 204 | val _ = print " * * *"; | 
| 
bd6582be1562
synchronize access to shared reference and print proper total
 blanchet parents: 
50563diff
changeset | 205 |     val _ = print ("Options: " ^ commas options);
 | 
| 50964 
2a990baa09af
use precomputed MaSh/MePo data whenever available
 blanchet parents: 
50953diff
changeset | 206 | val oks = Par_List.map solve_goal (tag_list 1 lines) | 
| 50620 | 207 | val n = length oks | 
| 50952 | 208 | val [mepo_ok, mash_isar_ok, mash_prover_ok, mesh_isar_ok, mesh_prover_ok, | 
| 209 | isar_ok] = | |
| 51028 | 210 | if n = 0 then zeros else map Integer.sum (map_transpose I oks) | 
| 48241 | 211 | in | 
| 212 | ["Successes (of " ^ string_of_int n ^ " goals)", | |
| 48379 
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
 blanchet parents: 
48378diff
changeset | 213 | total_of MePoN mepo_ok n, | 
| 50952 | 214 | total_of MaSh_IsarN mash_isar_ok n, | 
| 215 | total_of MaSh_ProverN mash_prover_ok n, | |
| 216 | total_of MeSh_IsarN mesh_isar_ok n, | |
| 217 | total_of MeSh_ProverN mesh_prover_ok n, | |
| 48379 
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
 blanchet parents: 
48378diff
changeset | 218 | total_of IsarN isar_ok n] | 
| 50437 | 219 | |> cat_lines |> print | 
| 48241 | 220 | end | 
| 48235 | 221 | |
| 48234 | 222 | end; |