author | blanchet |
Mon, 03 Feb 2014 10:14:18 +0100 | |
changeset 55266 | d9d31354834e |
parent 55264 | 43473497fb65 |
permissions | -rw-r--r-- |
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
54828
diff
changeset
|
1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_isar_try0.ML |
52592 | 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 |
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
54828
diff
changeset
|
5 |
Try replacing calls to "metis" with calls to other proof methods to speed up proofs, eliminate |
54766
6ac273f176cd
store alternative proof methods in Isar data structure
blanchet
parents:
54765
diff
changeset
|
6 |
dependencies, and repair broken proof steps. |
52592 | 7 |
*) |
8 |
||
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
54828
diff
changeset
|
9 |
signature SLEDGEHAMMER_ISAR_TRY0 = |
52592 | 10 |
sig |
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
54828
diff
changeset
|
11 |
type isar_proof = Sledgehammer_Isar_Proof.isar_proof |
55213 | 12 |
type isar_preplay_data = Sledgehammer_Isar_Preplay.isar_preplay_data |
52592 | 13 |
|
55260 | 14 |
val try0_isar_proof : Proof.context -> Time.time -> isar_preplay_data Unsynchronized.ref -> |
15 |
isar_proof -> isar_proof |
|
54504 | 16 |
end; |
52592 | 17 |
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
54828
diff
changeset
|
18 |
structure Sledgehammer_Isar_Try0 : SLEDGEHAMMER_ISAR_TRY0 = |
52592 | 19 |
struct |
20 |
||
54712 | 21 |
open Sledgehammer_Util |
54828 | 22 |
open Sledgehammer_Reconstructor |
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
54828
diff
changeset
|
23 |
open Sledgehammer_Isar_Proof |
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
54828
diff
changeset
|
24 |
open Sledgehammer_Isar_Preplay |
52592 | 25 |
|
55255 | 26 |
fun step_with_method meth (Prove (qs, xs, l, t, subproofs, (facts, _))) = |
27 |
Prove (qs, xs, l, t, subproofs, (facts, [meth])) |
|
54712 | 28 |
|
29 |
val slack = seconds 0.05 |
|
52592 | 30 |
|
55258 | 31 |
fun try0_step _ _ _ (step as Let _) = step |
55260 | 32 |
| try0_step ctxt preplay_timeout preplay_data |
55255 | 33 |
(step as Prove (_, _, l, _, _, (_, meths as meth :: _))) = |
54712 | 34 |
let |
35 |
val timeout = |
|
55266 | 36 |
(case Lazy.force (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth) of |
54828 | 37 |
Played time => Time.+ (time, slack) |
54826 | 38 |
| _ => preplay_timeout) |
52592 | 39 |
|
55255 | 40 |
fun try_method meth = |
55258 | 41 |
(case preplay_isar_step ctxt timeout meth step of |
55255 | 42 |
outcome as Played _ => SOME (meth, outcome) |
54827 | 43 |
| _ => NONE) |
54712 | 44 |
in |
55243 | 45 |
(* FIXME: create variant after success *) |
55255 | 46 |
(case Par_List.get_some try_method meths of |
47 |
SOME (meth', outcome) => |
|
55264 | 48 |
let val step' = step_with_method meth' step in |
49 |
(set_preplay_outcomes_of_isar_step ctxt timeout preplay_data step' |
|
50 |
[(meth', Lazy.value outcome)]; |
|
51 |
step') |
|
52 |
end |
|
54712 | 53 |
| NONE => step) |
54 |
end |
|
52592 | 55 |
|
55258 | 56 |
val try0_isar_proof = map_isar_steps ooo try0_step |
52592 | 57 |
|
54504 | 58 |
end; |