|
1 (* Title: HOL/Tools/Sledgehammer/sledgehammer_try0.ML |
|
2 Author: Jasmin Blanchette, TU Muenchen |
|
3 Author: Steffen Juilf Smolka, TU Muenchen |
|
4 |
|
5 Try replacing calls to metis with calls to other proof methods in order to |
|
6 speed up proofs, eliminate dependencies, and repair broken proof steps. |
|
7 *) |
|
8 |
|
9 signature SLEDGEHAMMER_TRY0 = |
|
10 sig |
|
11 type isar_proof = Sledgehammer_Proof.isar_proof |
|
12 type preplay_interface = Sledgehammer_Preplay.preplay_interface |
|
13 |
|
14 val try0 : Time.time -> preplay_interface -> isar_proof -> isar_proof |
|
15 end |
|
16 |
|
17 structure Sledgehammer_Try0 : SLEDGEHAMMER_TRY0 = |
|
18 struct |
|
19 |
|
20 open Sledgehammer_Proof |
|
21 open Sledgehammer_Preplay |
|
22 |
|
23 |
|
24 fun reachable_labels proof = |
|
25 let |
|
26 val refs_of_step = |
|
27 byline_of_step #> (fn SOME (By ((lfs, _), _)) => lfs | NONE => []) |
|
28 |
|
29 val union = fold (Ord_List.insert label_ord) |
|
30 |
|
31 fun do_proof proof reachable = |
|
32 let |
|
33 val (steps, concl) = split_last (steps_of_proof proof) |
|
34 val reachable = |
|
35 union (refs_of_step concl) reachable |
|
36 |> union (the_list (label_of_step concl)) |
|
37 in |
|
38 fold_rev do_step steps reachable |
|
39 end |
|
40 |
|
41 and do_step (Let _) reachable = reachable |
|
42 | do_step (Prove (_, _, l, _, subproofs, By ((lfs, _), _))) reachable = |
|
43 if Ord_List.member label_ord reachable l |
|
44 then fold do_proof subproofs (union lfs reachable) |
|
45 else reachable |
|
46 |
|
47 in |
|
48 do_proof proof [] |
|
49 end |
|
50 |
|
51 (** removes steps not referenced by the final proof step or any of its |
|
52 predecessors **) |
|
53 fun remove_unreachable_steps ({set_time, ...} : preplay_interface) proof = |
|
54 let |
|
55 |
|
56 val reachable = reachable_labels proof |
|
57 |
|
58 fun do_proof (Proof (fix, assms, steps)) = |
|
59 Proof (fix, assms, do_steps steps) |
|
60 |
|
61 and do_steps steps = |
|
62 uncurry (fold_rev do_step) (split_last steps ||> single) |
|
63 |
|
64 and do_step (step as Let _) accu = step :: accu |
|
65 | do_step (Prove (qs, xs, l, t, subproofs, by)) accu = |
|
66 if Ord_List.member label_ord reachable l |
|
67 then Prove (qs, xs, l, t, map do_proof subproofs, by) :: accu |
|
68 else (set_time l zero_preplay_time; accu) |
|
69 |
|
70 in |
|
71 do_proof proof |
|
72 end |
|
73 |
|
74 |
|
75 fun variants (step as Let _) = raise Fail "Sledgehammer_Try0: variants" |
|
76 | variants (Prove (qs, xs, l, t, subproofs, By (facts, method))) = |
|
77 let |
|
78 val methods = [SimpM, AutoM, FastforceM, ArithM] |
|
79 fun step method = Prove (qs, xs, l, t, subproofs, By (facts, method)) |
|
80 fun step_without_facts method = |
|
81 Prove (qs, xs, l, t, subproofs, By (no_facts, method)) |
|
82 in |
|
83 (* There seems to be a bias for methods earlier in the list, so we put |
|
84 the variants using no facts first. *) |
|
85 map step_without_facts methods @ map step methods |
|
86 end |
|
87 |
|
88 fun try0_step preplay_timeout preplay_interface (step as Let _) = step |
|
89 | try0_step preplay_timeout ({preplay_quietly, get_time, set_time, |
|
90 set_preplay_fail, ...} : preplay_interface) |
|
91 (step as Prove (_, _, l, _, _, _)) = |
|
92 let |
|
93 |
|
94 val (preplay_fail, timeout) = |
|
95 case get_time l of |
|
96 (true, _) => (true, preplay_timeout) |
|
97 | (false, t) => (false, t) |
|
98 |
|
99 fun try_variant variant = |
|
100 case preplay_quietly timeout variant of |
|
101 (true, _) => NONE |
|
102 | time as (false, _) => SOME (variant, time) |
|
103 |
|
104 in |
|
105 case Par_List.get_some try_variant (variants step) of |
|
106 SOME (step, time) => (set_time l time; step) |
|
107 | NONE => (if preplay_fail then set_preplay_fail true else (); step) |
|
108 end |
|
109 |
|
110 fun try0 preplay_timeout |
|
111 (preplay_interface as {set_preplay_fail, ...} : preplay_interface) proof = |
|
112 (set_preplay_fail false; (* reset preplay fail *) |
|
113 map_isar_steps (try0_step preplay_timeout preplay_interface) proof |
|
114 |> remove_unreachable_steps preplay_interface) |
|
115 |
|
116 end |