| author | wenzelm |
| Thu, 29 Aug 2013 15:53:56 +0200 | |
| changeset 53281 | 251e1a2aa792 |
| parent 52629 | d6f2a7c196f7 |
| child 54504 | 096f7d452164 |
| permissions | -rw-r--r-- |
| 52592 | 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 |
||
| 52612 | 23 |
fun variants (Let _) = raise Fail "Sledgehammer_Try0: variants" |
24 |
| variants (Prove (qs, xs, l, t, subproofs, By (facts, _))) = |
|
| 52592 | 25 |
let |
| 52629 | 26 |
val methods = [SimpM, FastforceM, AutoM, ArithM, ForceM, BlastM] |
| 52592 | 27 |
fun step method = Prove (qs, xs, l, t, subproofs, By (facts, method)) |
| 52612 | 28 |
(*fun step_without_facts method = |
29 |
Prove (qs, xs, l, t, subproofs, By (no_facts, method))*) |
|
| 52592 | 30 |
in |
|
52611
831f7479c74f
minimize dependencies (used facts) of Isar proof steps; remove unreferenced steps
smolkas
parents:
52592
diff
changeset
|
31 |
(* FIXME *) |
| 52592 | 32 |
(* There seems to be a bias for methods earlier in the list, so we put |
33 |
the variants using no facts first. *) |
|
|
52611
831f7479c74f
minimize dependencies (used facts) of Isar proof steps; remove unreferenced steps
smolkas
parents:
52592
diff
changeset
|
34 |
(*map step_without_facts methods @*) map step methods |
| 52592 | 35 |
end |
36 |
||
| 52612 | 37 |
fun try0_step _ _ (step as Let _) = step |
| 52626 | 38 |
| try0_step preplay_timeout ({preplay_quietly, get_preplay_time,
|
39 |
set_preplay_time, ...} : preplay_interface) |
|
| 52592 | 40 |
(step as Prove (_, _, l, _, _, _)) = |
41 |
let |
|
42 |
||
| 52626 | 43 |
val timeout = |
44 |
case get_preplay_time l of |
|
45 |
(true, _) => preplay_timeout |
|
46 |
| (false, t) => t |
|
| 52592 | 47 |
|
48 |
fun try_variant variant = |
|
49 |
case preplay_quietly timeout variant of |
|
50 |
(true, _) => NONE |
|
51 |
| time as (false, _) => SOME (variant, time) |
|
52 |
||
53 |
in |
|
54 |
case Par_List.get_some try_variant (variants step) of |
|
| 52626 | 55 |
SOME (step, time) => (set_preplay_time l time; step) |
56 |
| NONE => step |
|
| 52592 | 57 |
end |
58 |
||
| 52626 | 59 |
fun try0 preplay_timeout preplay_interface proof = |
60 |
map_isar_steps (try0_step preplay_timeout preplay_interface) proof |
|
| 52592 | 61 |
|
62 |
end |