| author | nipkow | 
| Sun, 28 Oct 2018 11:00:00 +0100 | |
| changeset 69200 | f2bb47056d8f | 
| parent 63806 | c54a53ef1873 | 
| permissions | -rw-r--r-- | 
| 47847 | 1  | 
(* Title: HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML  | 
| 38892 | 2  | 
Author: Jasmin Blanchette, TU Munich  | 
3  | 
*)  | 
|
4  | 
||
5  | 
structure Mirabelle_Sledgehammer_Filter : MIRABELLE_ACTION =  | 
|
6  | 
struct  | 
|
7  | 
||
| 
40070
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
8  | 
fun get args name default_value =  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
9  | 
case AList.lookup (op =) args name of  | 
| 63806 | 10  | 
SOME value => Value.parse_real value  | 
| 
40070
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
11  | 
| NONE => default_value  | 
| 
38992
 
542474156c66
introduce fudge factors to deal with "theory const"
 
blanchet 
parents: 
38988 
diff
changeset
 | 
12  | 
|
| 
40070
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
13  | 
fun extract_relevance_fudge args  | 
| 55205 | 14  | 
      {local_const_multiplier, worse_irrel_freq, higher_order_irrel_weight, abs_rel_weight,
 | 
15  | 
abs_irrel_weight, theory_const_rel_weight, theory_const_irrel_weight,  | 
|
16  | 
chained_const_irrel_weight, intro_bonus, elim_bonus, simp_bonus, local_bonus, assum_bonus,  | 
|
17  | 
chained_bonus, max_imperfect, max_imperfect_exp, threshold_divisor, ridiculous_threshold} =  | 
|
18  | 
  {local_const_multiplier = get args "local_const_multiplier" local_const_multiplier,
 | 
|
| 
41159
 
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
 
blanchet 
parents: 
41138 
diff
changeset
 | 
19  | 
worse_irrel_freq = get args "worse_irrel_freq" worse_irrel_freq,  | 
| 55205 | 20  | 
higher_order_irrel_weight = get args "higher_order_irrel_weight" higher_order_irrel_weight,  | 
| 
40070
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
21  | 
abs_rel_weight = get args "abs_rel_weight" abs_rel_weight,  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
22  | 
abs_irrel_weight = get args "abs_irrel_weight" abs_irrel_weight,  | 
| 55205 | 23  | 
theory_const_rel_weight = get args "theory_const_rel_weight" theory_const_rel_weight,  | 
24  | 
theory_const_irrel_weight = get args "theory_const_irrel_weight" theory_const_irrel_weight,  | 
|
25  | 
chained_const_irrel_weight = get args "chained_const_irrel_weight" chained_const_irrel_weight,  | 
|
| 
40070
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
26  | 
intro_bonus = get args "intro_bonus" intro_bonus,  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
27  | 
elim_bonus = get args "elim_bonus" elim_bonus,  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
28  | 
simp_bonus = get args "simp_bonus" simp_bonus,  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
29  | 
local_bonus = get args "local_bonus" local_bonus,  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
30  | 
assum_bonus = get args "assum_bonus" assum_bonus,  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
31  | 
chained_bonus = get args "chained_bonus" chained_bonus,  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
32  | 
max_imperfect = get args "max_imperfect" max_imperfect,  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
33  | 
max_imperfect_exp = get args "max_imperfect_exp" max_imperfect_exp,  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
34  | 
threshold_divisor = get args "threshold_divisor" threshold_divisor,  | 
| 
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
35  | 
ridiculous_threshold = get args "ridiculous_threshold" ridiculous_threshold}  | 
| 38902 | 36  | 
|
| 38892 | 37  | 
structure Prooftab =  | 
| 41407 | 38  | 
Table(type key = int * int val ord = prod_ord int_ord int_ord)  | 
| 38892 | 39  | 
|
| 41407 | 40  | 
val proof_table = Unsynchronized.ref (Prooftab.empty: string list list Prooftab.table)  | 
| 38892 | 41  | 
|
| 
38896
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
42  | 
val num_successes = Unsynchronized.ref ([] : (int * int) list)  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
43  | 
val num_failures = Unsynchronized.ref ([] : (int * int) list)  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
44  | 
val num_found_proofs = Unsynchronized.ref ([] : (int * int) list)  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
45  | 
val num_lost_proofs = Unsynchronized.ref ([] : (int * int) list)  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
46  | 
val num_found_facts = Unsynchronized.ref ([] : (int * int) list)  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
47  | 
val num_lost_facts = Unsynchronized.ref ([] : (int * int) list)  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
48  | 
|
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
49  | 
fun get id c = the_default 0 (AList.lookup (op =) (!c) id)  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
50  | 
fun add id c n =  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
51  | 
c := (case AList.lookup (op =) (!c) id of  | 
| 55205 | 52  | 
SOME m => AList.update (op =) (id, m + n) (!c)  | 
53  | 
| NONE => (id, n) :: !c)  | 
|
| 38894 | 54  | 
|
55  | 
fun init proof_file _ thy =  | 
|
| 38892 | 56  | 
let  | 
57  | 
fun do_line line =  | 
|
| 55205 | 58  | 
(case line |> space_explode ":" of  | 
| 
43710
 
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
 
wenzelm 
parents: 
43351 
diff
changeset
 | 
59  | 
[line_num, offset, proof] =>  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
60  | 
SOME (apply2 (the o Int.fromString) (line_num, offset),  | 
| 38892 | 61  | 
proof |> space_explode " " |> filter_out (curry (op =) ""))  | 
| 55205 | 62  | 
| _ => NONE)  | 
| 38894 | 63  | 
val proofs = File.read (Path.explode proof_file)  | 
| 38892 | 64  | 
val proof_tab =  | 
65  | 
proofs |> space_explode "\n"  | 
|
66  | 
|> map_filter do_line  | 
|
67  | 
|> AList.coalesce (op =)  | 
|
68  | 
|> Prooftab.make  | 
|
69  | 
in proof_table := proof_tab; thy end  | 
|
70  | 
||
| 38894 | 71  | 
fun percentage a b = if b = 0 then "N/A" else string_of_int (a * 100 div b)  | 
72  | 
fun percentage_alt a b = percentage a (a + b)  | 
|
73  | 
||
| 
38896
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
74  | 
fun done id ({log, ...} : Mirabelle.done_args) =
 | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
75  | 
if get id num_successes + get id num_failures > 0 then  | 
| 38897 | 76  | 
(log "";  | 
| 55205 | 77  | 
     log ("Number of overall successes: " ^ string_of_int (get id num_successes));
 | 
| 38897 | 78  | 
     log ("Number of overall failures: " ^ string_of_int (get id num_failures));
 | 
| 38894 | 79  | 
     log ("Overall success rate: " ^
 | 
| 
38896
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
80  | 
percentage_alt (get id num_successes) (get id num_failures) ^ "%");  | 
| 38897 | 81  | 
     log ("Number of found proofs: " ^ string_of_int (get id num_found_proofs));
 | 
82  | 
     log ("Number of lost proofs: " ^ string_of_int (get id num_lost_proofs));
 | 
|
| 38894 | 83  | 
     log ("Proof found rate: " ^
 | 
| 55205 | 84  | 
percentage_alt (get id num_found_proofs) (get id num_lost_proofs) ^ "%");  | 
| 38897 | 85  | 
     log ("Number of found facts: " ^ string_of_int (get id num_found_facts));
 | 
86  | 
     log ("Number of lost facts: " ^ string_of_int (get id num_lost_facts));
 | 
|
| 38894 | 87  | 
     log ("Fact found rate: " ^
 | 
| 55205 | 88  | 
percentage_alt (get id num_found_facts) (get id num_lost_facts) ^ "%"))  | 
| 38894 | 89  | 
else  | 
90  | 
()  | 
|
| 38892 | 91  | 
|
| 57154 | 92  | 
val default_prover = ATP_Proof.eN (* arbitrary ATP *)  | 
| 38892 | 93  | 
|
| 38897 | 94  | 
fun with_index (i, s) = s ^ "@" ^ string_of_int i  | 
95  | 
||
| 
38896
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
96  | 
fun action args id ({pre, pos, log, ...} : Mirabelle.run_args) =
 | 
| 
43710
 
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
 
wenzelm 
parents: 
43351 
diff
changeset
 | 
97  | 
case (Position.line_of pos, Position.offset_of pos) of  | 
| 
 
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
 
wenzelm 
parents: 
43351 
diff
changeset
 | 
98  | 
(SOME line_num, SOME offset) =>  | 
| 
 
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
 
wenzelm 
parents: 
43351 
diff
changeset
 | 
99  | 
(case Prooftab.lookup (!proof_table) (line_num, offset) of  | 
| 38892 | 100  | 
SOME proofs =>  | 
101  | 
let  | 
|
| 55205 | 102  | 
val thy = Proof.theory_of pre  | 
| 
43351
 
b19d95b4d736
compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
 
blanchet 
parents: 
43088 
diff
changeset
 | 
103  | 
         val {context = ctxt, facts = chained_ths, goal} = Proof.goal pre
 | 
| 54095 | 104  | 
val prover = AList.lookup (op =) args "prover" |> the_default default_prover  | 
| 55205 | 105  | 
         val params as {max_facts, ...} = Sledgehammer_Commands.default_params thy args
 | 
106  | 
val default_max_facts =  | 
|
107  | 
Sledgehammer_Prover_Minimize.default_max_facts_of_prover ctxt prover  | 
|
| 
40070
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
108  | 
val relevance_fudge =  | 
| 54095 | 109  | 
extract_relevance_fudge args Sledgehammer_MePo.default_relevance_fudge  | 
| 38892 | 110  | 
val subgoal = 1  | 
| 
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: 
52125 
diff
changeset
 | 
111  | 
val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal subgoal ctxt  | 
| 55212 | 112  | 
val ho_atp = Sledgehammer_Prover_ATP.is_ho_atp ctxt prover  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58921 
diff
changeset
 | 
113  | 
val keywords = Thy_Header.get_keywords' ctxt  | 
| 48299 | 114  | 
val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt  | 
| 38892 | 115  | 
val facts =  | 
| 48289 | 116  | 
Sledgehammer_Fact.nearly_all_facts ctxt ho_atp  | 
| 
58921
 
ffdafc99f67b
prefer explicit Keyword.keywords (cf. 82a71046dce8);
 
wenzelm 
parents: 
57154 
diff
changeset
 | 
117  | 
Sledgehammer_Fact.no_fact_override keywords css_table chained_ths  | 
| 48299 | 118  | 
hyp_ts concl_t  | 
| 
54143
 
18def1c73c79
make sure add: doesn't add duplicates, and works for [no_atp] facts
 
blanchet 
parents: 
54128 
diff
changeset
 | 
119  | 
|> Sledgehammer_Fact.drop_duplicate_facts  | 
| 48406 | 120  | 
|> Sledgehammer_MePo.mepo_suggested_facts ctxt params  | 
| 54095 | 121  | 
(the_default default_max_facts max_facts) (SOME relevance_fudge) hyp_ts concl_t  | 
| 
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: 
48406 
diff
changeset
 | 
122  | 
|> map (fst o fst)  | 
| 38894 | 123  | 
val (found_facts, lost_facts) =  | 
| 41455 | 124  | 
flat proofs |> sort_distinct string_ord  | 
| 38897 | 125  | 
|> map (fn fact => (find_index (curry (op =) fact) facts, fact))  | 
126  | 
|> List.partition (curry (op <=) 0 o fst)  | 
|
127  | 
|>> sort (prod_ord int_ord string_ord) ||> map snd  | 
|
| 38892 | 128  | 
val found_proofs = filter (forall (member (op =) facts)) proofs  | 
| 38894 | 129  | 
val n = length found_proofs  | 
| 38892 | 130  | 
val _ =  | 
| 38894 | 131  | 
if n = 0 then  | 
| 
38896
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
132  | 
(add id num_failures 1; log "Failure")  | 
| 38894 | 133  | 
else  | 
| 
38896
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
134  | 
(add id num_successes 1;  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
135  | 
add id num_found_proofs n;  | 
| 38897 | 136  | 
              log ("Success (" ^ string_of_int n ^ " of " ^
 | 
137  | 
string_of_int (length proofs) ^ " proofs)"))  | 
|
| 
38896
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
138  | 
val _ = add id num_lost_proofs (length proofs - n)  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
139  | 
val _ = add id num_found_facts (length found_facts)  | 
| 
 
b36ab8860748
allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
 
blanchet 
parents: 
38894 
diff
changeset
 | 
140  | 
val _ = add id num_lost_facts (length lost_facts)  | 
| 38902 | 141  | 
val _ =  | 
142  | 
if null found_facts then  | 
|
143  | 
()  | 
|
144  | 
else  | 
|
145  | 
let  | 
|
146  | 
val found_weight =  | 
|
| 55205 | 147  | 
Real.fromInt (fold (fn (n, _) => Integer.add (n * n)) found_facts 0)  | 
| 38902 | 148  | 
/ Real.fromInt (length found_facts)  | 
149  | 
|> Math.sqrt |> Real.ceil  | 
|
150  | 
in  | 
|
151  | 
               log ("Found facts (among " ^ string_of_int (length facts) ^
 | 
|
152  | 
", weight " ^ string_of_int found_weight ^ "): " ^  | 
|
153  | 
commas (map with_index found_facts))  | 
|
154  | 
end  | 
|
155  | 
val _ = if null lost_facts then  | 
|
| 38897 | 156  | 
()  | 
157  | 
else  | 
|
| 38902 | 158  | 
                   log ("Lost facts (among " ^ string_of_int (length facts) ^
 | 
159  | 
"): " ^ commas lost_facts)  | 
|
| 38892 | 160  | 
in () end  | 
| 38894 | 161  | 
| NONE => log "No known proof")  | 
| 38892 | 162  | 
| _ => ()  | 
163  | 
||
| 38894 | 164  | 
val proof_fileK = "proof_file"  | 
165  | 
||
166  | 
fun invoke args =  | 
|
167  | 
let  | 
|
| 55205 | 168  | 
val (pf_args, other_args) = args |> List.partition (curry (op =) proof_fileK o fst)  | 
169  | 
val proof_file =  | 
|
170  | 
(case pf_args of  | 
|
171  | 
[] => error "No \"proof_file\" specified"  | 
|
172  | 
| (_, s) :: _ => s)  | 
|
| 38894 | 173  | 
in Mirabelle.register (init proof_file, action other_args, done) end  | 
| 38892 | 174  | 
|
175  | 
end;  | 
|
| 38894 | 176  | 
|
177  | 
(* Workaround to keep the "mirabelle.pl" script happy *)  | 
|
178  | 
structure Mirabelle_Sledgehammer_filter = Mirabelle_Sledgehammer_Filter;  |