| author | blanchet |
| Tue, 20 May 2014 16:39:13 +0200 | |
| changeset 57019 | f013e3a830c3 |
| parent 56985 | 82c83978fbd9 |
| child 57054 | fed0329ea8e2 |
| permissions | -rw-r--r-- |
| 55287 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML |
| 52555 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
3 |
Author: Steffen Juilf Smolka, TU Muenchen |
|
4 |
||
5 |
Reconstructors. |
|
6 |
*) |
|
7 |
||
| 55287 | 8 |
signature SLEDGEHAMMER_PROOF_METHODS = |
| 52555 | 9 |
sig |
10 |
type stature = ATP_Problem_Generate.stature |
|
11 |
||
| 55285 | 12 |
datatype proof_method = |
13 |
Metis_Method of string option * string option | |
|
14 |
Meson_Method | |
|
| 56081 | 15 |
SMT2_Method | |
|
56852
b38c5b9cf590
added 'satx' to Sledgehammer's portfolio (cf. 'isar_try0')
blanchet
parents:
56093
diff
changeset
|
16 |
SATx_Method | |
|
55323
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents:
55315
diff
changeset
|
17 |
Blast_Method | |
| 55285 | 18 |
Simp_Method | |
19 |
Simp_Size_Method | |
|
20 |
Auto_Method | |
|
21 |
Fastforce_Method | |
|
22 |
Force_Method | |
|
|
55323
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents:
55315
diff
changeset
|
23 |
Linarith_Method | |
|
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents:
55315
diff
changeset
|
24 |
Presburger_Method | |
| 55285 | 25 |
Algebra_Method |
| 52555 | 26 |
|
| 54824 | 27 |
datatype play_outcome = |
28 |
Played of Time.time | |
|
29 |
Play_Timed_Out of Time.time | |
|
| 56093 | 30 |
Play_Failed |
| 52555 | 31 |
|
32 |
type minimize_command = string list -> string |
|
| 54824 | 33 |
type one_line_params = |
| 55285 | 34 |
(proof_method * play_outcome) * string * (string * stature) list * minimize_command * int * int |
| 52555 | 35 |
|
|
56985
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
blanchet
parents:
56983
diff
changeset
|
36 |
val string_of_proof_method : Proof.context -> string list -> proof_method -> string |
|
55452
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
blanchet
parents:
55451
diff
changeset
|
37 |
val tac_of_proof_method : Proof.context -> bool -> thm list * thm list -> proof_method -> int -> |
|
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
blanchet
parents:
55451
diff
changeset
|
38 |
tactic |
| 55211 | 39 |
val string_of_play_outcome : play_outcome -> string |
| 55269 | 40 |
val play_outcome_ord : play_outcome * play_outcome -> order |
|
56985
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
blanchet
parents:
56983
diff
changeset
|
41 |
val one_line_proof_text : Proof.context -> int -> one_line_params -> string |
| 54495 | 42 |
end; |
| 52555 | 43 |
|
| 55287 | 44 |
structure Sledgehammer_Proof_Methods : SLEDGEHAMMER_PROOF_METHODS = |
| 52555 | 45 |
struct |
46 |
||
| 54828 | 47 |
open ATP_Util |
| 52555 | 48 |
open ATP_Problem_Generate |
| 55211 | 49 |
open ATP_Proof_Reconstruct |
| 52555 | 50 |
|
| 55285 | 51 |
datatype proof_method = |
52 |
Metis_Method of string option * string option | |
|
53 |
Meson_Method | |
|
| 56081 | 54 |
SMT2_Method | |
|
56852
b38c5b9cf590
added 'satx' to Sledgehammer's portfolio (cf. 'isar_try0')
blanchet
parents:
56093
diff
changeset
|
55 |
SATx_Method | |
|
55323
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents:
55315
diff
changeset
|
56 |
Blast_Method | |
| 55285 | 57 |
Simp_Method | |
58 |
Simp_Size_Method | |
|
59 |
Auto_Method | |
|
60 |
Fastforce_Method | |
|
61 |
Force_Method | |
|
|
55323
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents:
55315
diff
changeset
|
62 |
Linarith_Method | |
|
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents:
55315
diff
changeset
|
63 |
Presburger_Method | |
| 55285 | 64 |
Algebra_Method |
| 52555 | 65 |
|
| 54824 | 66 |
datatype play_outcome = |
67 |
Played of Time.time | |
|
68 |
Play_Timed_Out of Time.time | |
|
| 56093 | 69 |
Play_Failed |
| 52555 | 70 |
|
| 55211 | 71 |
type minimize_command = string list -> string |
72 |
type one_line_params = |
|
| 55285 | 73 |
(proof_method * play_outcome) * string * (string * stature) list * minimize_command * int * int |
| 55211 | 74 |
|
|
56983
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
75 |
fun maybe_paren s = s |> not (Symbol_Pos.is_identifier s) ? enclose "(" ")"
|
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
76 |
|
|
56985
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
blanchet
parents:
56983
diff
changeset
|
77 |
fun string_of_proof_method ctxt ss meth = |
|
56983
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
78 |
let |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
79 |
val meth_s = |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
80 |
(case meth of |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
81 |
Metis_Method (NONE, NONE) => "metis" |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
82 |
| Metis_Method (type_enc_opt, lam_trans_opt) => |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
83 |
"metis (" ^ commas (map_filter I [type_enc_opt, lam_trans_opt]) ^ ")"
|
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
84 |
| Meson_Method => "meson" |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
85 |
| SMT2_Method => "smt2" |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
86 |
| SATx_Method => "satx" |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
87 |
| Blast_Method => "blast" |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
88 |
| Simp_Method => if null ss then "simp" else "simp add:" |
|
56985
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
blanchet
parents:
56983
diff
changeset
|
89 |
| Simp_Size_Method => "simp add: " ^ short_thm_name ctxt @{thm size_ne_size_imp_ne}
|
|
56983
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
90 |
| Auto_Method => "auto" |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
91 |
| Fastforce_Method => "fastforce" |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
92 |
| Force_Method => "force" |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
93 |
| Linarith_Method => "linarith" |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
94 |
| Presburger_Method => "presburger" |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
95 |
| Algebra_Method => "algebra") |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
96 |
in |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
97 |
maybe_paren (space_implode " " (meth_s :: ss)) |
|
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
98 |
end |
| 55285 | 99 |
|
|
56965
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
100 |
fun tac_of_proof_method ctxt debug (local_facts, global_facts) meth = |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
101 |
Method.insert_tac local_facts THEN' |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
102 |
(case meth of |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
103 |
Metis_Method (type_enc_opt, lam_trans_opt) => |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
104 |
Metis_Tactic.metis_tac [type_enc_opt |> the_default (hd partial_type_encs)] |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
105 |
(lam_trans_opt |> the_default default_metis_lam_trans) ctxt global_facts |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
106 |
| Meson_Method => Meson_Tactic.meson_general_tac ctxt global_facts |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
107 |
| SMT2_Method => SMT2_Solver.smt2_tac ctxt global_facts |
|
56983
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents:
56965
diff
changeset
|
108 |
| Simp_Method => Simplifier.asm_full_simp_tac (ctxt addsimps global_facts) |
|
56965
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
109 |
| _ => |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
110 |
Method.insert_tac global_facts THEN' |
| 55285 | 111 |
(case meth of |
|
56965
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
112 |
SATx_Method => SAT.satx_tac ctxt |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
113 |
| Blast_Method => blast_tac ctxt |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
114 |
| Simp_Size_Method => |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
115 |
Simplifier.asm_full_simp_tac (Simplifier.add_simp @{thm size_ne_size_imp_ne} ctxt)
|
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
116 |
| Auto_Method => K (Clasimp.auto_tac ctxt) |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
117 |
| Fastforce_Method => Clasimp.fast_force_tac ctxt |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
118 |
| Force_Method => Clasimp.force_tac ctxt |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
119 |
| Linarith_Method => Lin_Arith.tac ctxt |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
120 |
| Presburger_Method => Cooper.tac true [] [] ctxt |
|
222f46a4dbec
new approach to silence proof methods, to avoid weird theory/context mismatches
blanchet
parents:
56951
diff
changeset
|
121 |
| Algebra_Method => Groebner.algebra_tac [] [] ctxt)) |
| 55211 | 122 |
|
| 54828 | 123 |
fun string_of_play_outcome (Played time) = string_of_ext_time (false, time) |
| 56093 | 124 |
| string_of_play_outcome (Play_Timed_Out time) = |
125 |
if time = Time.zeroTime then "" else string_of_ext_time (true, time) ^ ", timed out" |
|
| 54828 | 126 |
| string_of_play_outcome Play_Failed = "failed" |
127 |
||
| 55269 | 128 |
fun play_outcome_ord (Played time1, Played time2) = |
129 |
int_ord (pairself Time.toMilliseconds (time1, time2)) |
|
130 |
| play_outcome_ord (Played _, _) = LESS |
|
131 |
| play_outcome_ord (_, Played _) = GREATER |
|
132 |
| play_outcome_ord (Play_Timed_Out time1, Play_Timed_Out time2) = |
|
133 |
int_ord (pairself Time.toMilliseconds (time1, time2)) |
|
134 |
| play_outcome_ord (Play_Timed_Out _, _) = LESS |
|
135 |
| play_outcome_ord (_, Play_Timed_Out _) = GREATER |
|
136 |
| play_outcome_ord (Play_Failed, Play_Failed) = EQUAL |
|
137 |
||
| 55211 | 138 |
fun apply_on_subgoal _ 1 = "by " |
139 |
| apply_on_subgoal 1 _ = "apply " |
|
140 |
| apply_on_subgoal i n = |
|
141 |
"prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal 1 n |
|
142 |
||
| 55285 | 143 |
(* FIXME *) |
|
56985
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
blanchet
parents:
56983
diff
changeset
|
144 |
fun proof_method_command ctxt meth i n _(*used_chaineds*) _(*num_chained*) ss = |
|
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
blanchet
parents:
56983
diff
changeset
|
145 |
apply_on_subgoal i n ^ string_of_proof_method ctxt ss meth |
| 55211 | 146 |
|
| 56093 | 147 |
fun try_command_line banner play command = |
148 |
let val s = string_of_play_outcome play in |
|
149 |
banner ^ ": " ^ Active.sendback_markup [Markup.padding_command] command ^ |
|
150 |
(s |> s <> "" ? enclose " (" ")") ^ "."
|
|
151 |
end |
|
| 52555 | 152 |
|
| 55211 | 153 |
fun minimize_line _ [] = "" |
154 |
| minimize_line minimize_command ss = |
|
155 |
(case minimize_command ss of |
|
156 |
"" => "" |
|
157 |
| command => "\nTo minimize: " ^ Active.sendback_markup [Markup.padding_command] command ^ ".") |
|
158 |
||
159 |
fun split_used_facts facts = |
|
160 |
facts |
|
161 |
|> List.partition (fn (_, (sc, _)) => sc = Chained) |
|
162 |
|> pairself (sort_distinct (string_ord o pairself fst)) |
|
163 |
||
|
56985
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
blanchet
parents:
56983
diff
changeset
|
164 |
fun one_line_proof_text ctxt num_chained |
| 55285 | 165 |
((meth, play), banner, used_facts, minimize_command, subgoal, subgoal_count) = |
| 55211 | 166 |
let |
167 |
val (chained, extra) = split_used_facts used_facts |
|
168 |
||
169 |
val try_line = |
|
170 |
map fst extra |
|
|
56985
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
blanchet
parents:
56983
diff
changeset
|
171 |
|> proof_method_command ctxt meth subgoal subgoal_count (map fst chained) num_chained |
| 56093 | 172 |
|> (if play = Play_Failed then enclose "One-line proof reconstruction failed: " "." |
173 |
else try_command_line banner play) |
|
| 55211 | 174 |
in |
175 |
try_line ^ minimize_line minimize_command (map fst (extra @ chained)) |
|
176 |
end |
|
| 52555 | 177 |
|
| 54495 | 178 |
end; |