| author | wenzelm | 
| Sun, 23 Feb 2014 15:35:19 +0100 | |
| changeset 55688 | 767edb2c1e4e | 
| parent 55452 | 29ec8680e61f | 
| child 56081 | 72fad75baf7e | 
| 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 | | |
| 15 | SMT_Method | | |
| 55323 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55315diff
changeset | 16 | Blast_Method | | 
| 55285 | 17 | Simp_Method | | 
| 18 | Simp_Size_Method | | |
| 19 | Auto_Method | | |
| 20 | Fastforce_Method | | |
| 21 | Force_Method | | |
| 55323 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55315diff
changeset | 22 | Linarith_Method | | 
| 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55315diff
changeset | 23 | Presburger_Method | | 
| 55285 | 24 | Algebra_Method | 
| 52555 | 25 | |
| 54824 | 26 | datatype play_outcome = | 
| 27 | Played of Time.time | | |
| 28 | Play_Timed_Out of Time.time | | |
| 29 | Play_Failed | | |
| 30 | Not_Played | |
| 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 | |
| 55285 | 36 | val string_of_proof_method : 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: 
55451diff
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: 
55451diff
changeset | 38 | tactic | 
| 55211 | 39 | val string_of_play_outcome : play_outcome -> string | 
| 55269 | 40 | val play_outcome_ord : play_outcome * play_outcome -> order | 
| 55211 | 41 | val one_line_proof_text : 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 | | |
| 54 | SMT_Method | | |
| 55323 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55315diff
changeset | 55 | Blast_Method | | 
| 55285 | 56 | Simp_Method | | 
| 57 | Simp_Size_Method | | |
| 58 | Auto_Method | | |
| 59 | Fastforce_Method | | |
| 60 | Force_Method | | |
| 55323 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55315diff
changeset | 61 | Linarith_Method | | 
| 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55315diff
changeset | 62 | Presburger_Method | | 
| 55285 | 63 | Algebra_Method | 
| 52555 | 64 | |
| 54824 | 65 | datatype play_outcome = | 
| 66 | Played of Time.time | | |
| 67 | Play_Timed_Out of Time.time | | |
| 68 | Play_Failed | | |
| 69 | Not_Played | |
| 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 | |
| 55285 | 75 | fun string_of_proof_method meth = | 
| 76 | (case meth of | |
| 77 | Metis_Method (NONE, NONE) => "metis" | |
| 78 | | Metis_Method (type_enc_opt, lam_trans_opt) => | |
| 79 |     "metis (" ^ commas (map_filter I [type_enc_opt, lam_trans_opt]) ^ ")"
 | |
| 80 | | Meson_Method => "meson" | |
| 81 | | SMT_Method => "smt" | |
| 55323 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55315diff
changeset | 82 | | Blast_Method => "blast" | 
| 55285 | 83 | | Simp_Method => "simp" | 
| 84 | | Simp_Size_Method => "simp add: size_ne_size_imp_ne" | |
| 85 | | Auto_Method => "auto" | |
| 86 | | Fastforce_Method => "fastforce" | |
| 87 | | Force_Method => "force" | |
| 55323 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55315diff
changeset | 88 | | Linarith_Method => "linarith" | 
| 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55315diff
changeset | 89 | | Presburger_Method => "presburger" | 
| 55285 | 90 | | Algebra_Method => "algebra") | 
| 91 | ||
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 92 | (* Makes proof methods as silent as possible. The "set_visible" calls suppresses "Unification bound | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 93 | exceeded" warnings and the like. *) | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 94 | fun silence_proof_methods debug = | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 95 | Try0.silence_methods debug | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 96 | #> Config.put SMT_Config.verbose debug | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 97 | |
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 98 | fun tac_of_proof_method ctxt0 debug (local_facts, global_facts) meth = | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 99 | let val ctxt = silence_proof_methods debug ctxt0 in | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 100 | Method.insert_tac local_facts THEN' | 
| 55285 | 101 | (case meth of | 
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 102 | Metis_Method (type_enc_opt, lam_trans_opt) => | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 103 | Metis_Tactic.metis_tac [type_enc_opt |> the_default (hd partial_type_encs)] | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 104 | (lam_trans_opt |> the_default default_metis_lam_trans) ctxt global_facts | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 105 | | Meson_Method => Meson.meson_tac ctxt global_facts | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 106 | |
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 107 | | SMT_Method => SMT_Solver.smt_tac ctxt global_facts | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 108 | | _ => | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 109 | Method.insert_tac global_facts THEN' | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 110 | (case meth of | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 111 | Blast_Method => blast_tac ctxt | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 112 | | Simp_Method => Simplifier.asm_full_simp_tac ctxt | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 113 | | Simp_Size_Method => | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 114 |         Simplifier.asm_full_simp_tac (Simplifier.add_simp @{thm size_ne_size_imp_ne} ctxt)
 | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 115 | | Auto_Method => K (Clasimp.auto_tac ctxt) | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 116 | | Fastforce_Method => Clasimp.fast_force_tac ctxt | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 117 | | Force_Method => Clasimp.force_tac ctxt | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 118 | | Linarith_Method => Lin_Arith.tac ctxt | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 119 | | Presburger_Method => Cooper.tac true [] [] ctxt | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 120 | | Algebra_Method => Groebner.algebra_tac [] [] ctxt)) | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55451diff
changeset | 121 | end | 
| 55211 | 122 | |
| 54828 | 123 | fun string_of_play_outcome (Played time) = string_of_ext_time (false, time) | 
| 124 | | string_of_play_outcome (Play_Timed_Out time) = string_of_ext_time (true, time) ^ ", timed out" | |
| 125 | | string_of_play_outcome Play_Failed = "failed" | |
| 126 | | string_of_play_outcome _ = "unknown" | |
| 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 (Not_Played, Not_Played) = EQUAL | |
| 133 | | play_outcome_ord (Not_Played, _) = LESS | |
| 134 | | play_outcome_ord (_, Not_Played) = GREATER | |
| 135 | | play_outcome_ord (Play_Timed_Out time1, Play_Timed_Out time2) = | |
| 136 | int_ord (pairself Time.toMilliseconds (time1, time2)) | |
| 137 | | play_outcome_ord (Play_Timed_Out _, _) = LESS | |
| 138 | | play_outcome_ord (_, Play_Timed_Out _) = GREATER | |
| 139 | | play_outcome_ord (Play_Failed, Play_Failed) = EQUAL | |
| 140 | ||
| 55211 | 141 | (* FIXME: Various bugs, esp. with "unfolding" | 
| 142 | fun unusing_chained_facts _ 0 = "" | |
| 143 | | unusing_chained_facts used_chaineds num_chained = | |
| 144 | if length used_chaineds = num_chained then "" | |
| 145 | else if null used_chaineds then "(* using no facts *) " | |
| 146 | else "(* using only " ^ space_implode " " used_chaineds ^ " *) " | |
| 147 | *) | |
| 148 | ||
| 149 | fun apply_on_subgoal _ 1 = "by " | |
| 150 | | apply_on_subgoal 1 _ = "apply " | |
| 151 | | apply_on_subgoal i n = | |
| 152 | "prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal 1 n | |
| 153 | ||
| 154 | fun command_call name [] = | |
| 155 |     name |> not (Symbol_Pos.is_identifier name) ? enclose "(" ")"
 | |
| 156 |   | command_call name args = "(" ^ name ^ " " ^ space_implode " " args ^ ")"
 | |
| 157 | ||
| 55285 | 158 | (* FIXME *) | 
| 159 | fun proof_method_command meth i n used_chaineds num_chained ss = | |
| 55211 | 160 | (* unusing_chained_facts used_chaineds num_chained ^ *) | 
| 55285 | 161 | apply_on_subgoal i n ^ command_call (string_of_proof_method meth) ss | 
| 55211 | 162 | |
| 163 | fun show_time NONE = "" | |
| 164 |   | show_time (SOME ext_time) = " (" ^ string_of_ext_time ext_time ^ ")"
 | |
| 165 | ||
| 166 | fun try_command_line banner time command = | |
| 167 | banner ^ ": " ^ Active.sendback_markup [Markup.padding_command] command ^ show_time time ^ "." | |
| 52555 | 168 | |
| 55211 | 169 | fun minimize_line _ [] = "" | 
| 170 | | minimize_line minimize_command ss = | |
| 171 | (case minimize_command ss of | |
| 172 | "" => "" | |
| 173 | | command => "\nTo minimize: " ^ Active.sendback_markup [Markup.padding_command] command ^ ".") | |
| 174 | ||
| 175 | fun split_used_facts facts = | |
| 176 | facts | |
| 177 | |> List.partition (fn (_, (sc, _)) => sc = Chained) | |
| 178 | |> pairself (sort_distinct (string_ord o pairself fst)) | |
| 179 | ||
| 180 | fun one_line_proof_text num_chained | |
| 55285 | 181 | ((meth, play), banner, used_facts, minimize_command, subgoal, subgoal_count) = | 
| 55211 | 182 | let | 
| 183 | val (chained, extra) = split_used_facts used_facts | |
| 184 | ||
| 185 | val (failed, ext_time) = | |
| 186 | (case play of | |
| 187 | Played time => (false, (SOME (false, time))) | |
| 188 | | Play_Timed_Out time => (false, SOME (true, time)) | |
| 189 | | Play_Failed => (true, NONE) | |
| 190 | | Not_Played => (false, NONE)) | |
| 191 | ||
| 192 | val try_line = | |
| 193 | map fst extra | |
| 55285 | 194 | |> proof_method_command meth subgoal subgoal_count (map fst chained) num_chained | 
| 55451 | 195 | |> (if failed then enclose "One-line proof reconstruction failed: " "." | 
| 196 | else try_command_line banner ext_time) | |
| 55211 | 197 | in | 
| 198 | try_line ^ minimize_line minimize_command (map fst (extra @ chained)) | |
| 199 | end | |
| 52555 | 200 | |
| 54495 | 201 | end; |