author | blanchet |
Thu, 19 Dec 2013 18:39:54 +0100 | |
changeset 54827 | 14e2c7616209 |
parent 54826 | 79745ba60a5a |
child 54828 | b2271ad695db |
permissions | -rw-r--r-- |
52592 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_try0.ML |
2 |
Author: Steffen Juilf Smolka, TU Muenchen |
|
54766
6ac273f176cd
store alternative proof methods in Isar data structure
blanchet
parents:
54765
diff
changeset
|
3 |
Author: Jasmin Blanchette, TU Muenchen |
52592 | 4 |
|
54766
6ac273f176cd
store alternative proof methods in Isar data structure
blanchet
parents:
54765
diff
changeset
|
5 |
Try replacing calls to metis with calls to other proof methods to speed up proofs, eliminate |
6ac273f176cd
store alternative proof methods in Isar data structure
blanchet
parents:
54765
diff
changeset
|
6 |
dependencies, and repair broken proof steps. |
52592 | 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 |
|
54504 | 15 |
end; |
52592 | 16 |
|
17 |
structure Sledgehammer_Try0 : SLEDGEHAMMER_TRY0 = |
|
18 |
struct |
|
19 |
||
54712 | 20 |
open Sledgehammer_Util |
52592 | 21 |
open Sledgehammer_Proof |
22 |
open Sledgehammer_Preplay |
|
23 |
||
54767
81290fe85890
generalize method list further to list of list (clustering preferred methods together)
blanchet
parents:
54766
diff
changeset
|
24 |
fun variants_of_step (Prove (qs, xs, l, t, subproofs, (facts, methss))) = |
81290fe85890
generalize method list further to list of list (clustering preferred methods together)
blanchet
parents:
54766
diff
changeset
|
25 |
map (fn meth => Prove (qs, xs, l, t, subproofs, (facts, [[meth]]))) (tl (flat methss)) |
54766
6ac273f176cd
store alternative proof methods in Isar data structure
blanchet
parents:
54765
diff
changeset
|
26 |
| variants_of_step _ = raise Fail "Sledgehammer_Try0: variants_of_step" |
54712 | 27 |
|
28 |
val slack = seconds 0.05 |
|
52592 | 29 |
|
52612 | 30 |
fun try0_step _ _ (step as Let _) = step |
54712 | 31 |
| try0_step preplay_timeout |
54827 | 32 |
({get_preplay_result, set_preplay_result, preplay_quietly, ...} : preplay_interface) |
54712 | 33 |
(step as Prove (_, _, l, _, _, _)) = |
34 |
let |
|
35 |
val timeout = |
|
54826 | 36 |
(case get_preplay_result l of |
37 |
Preplay_Success (false, t) => Time.+ (t, slack) |
|
38 |
| _ => preplay_timeout) |
|
52592 | 39 |
|
54712 | 40 |
fun try_variant variant = |
41 |
(case preplay_quietly timeout variant of |
|
54827 | 42 |
result as Preplay_Success _ => SOME (variant, result) |
43 |
| _ => NONE) |
|
54712 | 44 |
in |
45 |
(case Par_List.get_some try_variant (variants_of_step step) of |
|
54827 | 46 |
SOME (step', result) => (set_preplay_result l result; step') |
54712 | 47 |
| NONE => step) |
48 |
end |
|
52592 | 49 |
|
54712 | 50 |
val try0 = map_isar_steps oo try0_step |
52592 | 51 |
|
54504 | 52 |
end; |