| author | paulson <lp15@cam.ac.uk> | 
| Mon, 27 Feb 2023 17:09:59 +0000 | |
| changeset 77406 | c2013f617a70 | 
| parent 77269 | bc43f86c9598 | 
| child 77418 | a8458f0df4ee | 
| permissions | -rw-r--r-- | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55201diff
changeset | 1 | (* Title: HOL/Tools/Sledgehammer/sledgehammer.ML | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 2 | Author: Fabian Immler, TU Muenchen | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 3 | Author: Makarius | 
| 35969 | 4 | Author: Jasmin Blanchette, TU Muenchen | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 5 | |
| 38021 
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
 blanchet parents: 
38020diff
changeset | 6 | Sledgehammer's heart. | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 7 | *) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 8 | |
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55201diff
changeset | 9 | signature SLEDGEHAMMER = | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 10 | sig | 
| 60612 
79d71bfea310
removed chained facts from preplaying -- and careful about extra chained facts when removing 'proof -' and 'qed' from one-line Isar proofs
 blanchet parents: 
60549diff
changeset | 11 | type stature = ATP_Problem_Generate.stature | 
| 51008 | 12 | type fact = Sledgehammer_Fact.fact | 
| 48292 | 13 | type fact_override = Sledgehammer_Fact.fact_override | 
| 57755 | 14 | type proof_method = Sledgehammer_Proof_Methods.proof_method | 
| 15 | type play_outcome = Sledgehammer_Proof_Methods.play_outcome | |
| 55201 | 16 | type mode = Sledgehammer_Prover.mode | 
| 17 | type params = Sledgehammer_Prover.params | |
| 74951 
0b6f795d3b78
proper filtering inf induction rules in Mirabelle
 desharna parents: 
74950diff
changeset | 18 | type induction_rules = Sledgehammer_Prover.induction_rules | 
| 74952 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 19 | type prover_problem = Sledgehammer_Prover.prover_problem | 
| 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 20 | type prover_result = Sledgehammer_Prover.prover_result | 
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 21 | |
| 76524 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 22 | type preplay_result = proof_method * (play_outcome * (string * stature) list) | 
| 75372 | 23 | |
| 75020 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 24 | datatype sledgehammer_outcome = | 
| 75372 | 25 | SH_Some of prover_result * preplay_result list | 
| 75020 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 26 | | SH_Unknown | 
| 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 27 | | SH_Timeout | 
| 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 28 | | SH_None | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 29 | |
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 30 | val short_string_of_sledgehammer_outcome : sledgehammer_outcome -> string | 
| 51010 | 31 | val string_of_factss : (string * fact list) list -> string | 
| 54816 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 blanchet parents: 
54799diff
changeset | 32 | val run_sledgehammer : params -> mode -> (string -> unit) option -> int -> fact_override -> | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 33 | Proof.state -> bool * (sledgehammer_outcome * string) | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 34 | end; | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 35 | |
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55201diff
changeset | 36 | structure Sledgehammer : SLEDGEHAMMER = | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 37 | struct | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 38 | |
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43064diff
changeset | 39 | open ATP_Util | 
| 75029 | 40 | open ATP_Problem | 
| 55212 | 41 | open ATP_Proof | 
| 46320 | 42 | open ATP_Problem_Generate | 
| 38023 | 43 | open Sledgehammer_Util | 
| 48250 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 blanchet parents: 
47904diff
changeset | 44 | open Sledgehammer_Fact | 
| 55287 | 45 | open Sledgehammer_Proof_Methods | 
| 57734 
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
 blanchet parents: 
57557diff
changeset | 46 | open Sledgehammer_Isar_Proof | 
| 
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
 blanchet parents: 
57557diff
changeset | 47 | open Sledgehammer_Isar_Preplay | 
| 57740 | 48 | open Sledgehammer_Isar_Minimize | 
| 75029 | 49 | open Sledgehammer_ATP_Systems | 
| 55201 | 50 | open Sledgehammer_Prover | 
| 55212 | 51 | open Sledgehammer_Prover_ATP | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55201diff
changeset | 52 | open Sledgehammer_Prover_Minimize | 
| 48381 | 53 | open Sledgehammer_MaSh | 
| 40072 
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
 blanchet parents: 
40071diff
changeset | 54 | |
| 76524 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 55 | type preplay_result = proof_method * (play_outcome * (string * stature) list) | 
| 75372 | 56 | |
| 75020 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 57 | datatype sledgehammer_outcome = | 
| 75372 | 58 | SH_Some of prover_result * preplay_result list | 
| 75020 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 59 | | SH_Unknown | 
| 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 60 | | SH_Timeout | 
| 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 61 | | SH_None | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 62 | |
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 63 | fun short_string_of_sledgehammer_outcome (SH_Some _) = "some" | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 64 | | short_string_of_sledgehammer_outcome SH_Unknown = "unknown" | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 65 | | short_string_of_sledgehammer_outcome SH_Timeout = "timeout" | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 66 | | short_string_of_sledgehammer_outcome SH_None = "none" | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 67 | |
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 68 | fun alternative f (SOME x) (SOME y) = SOME (f (x, y)) | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 69 | | alternative _ (x as SOME _) NONE = x | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 70 | | alternative _ NONE (y as SOME _) = y | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 71 | | alternative _ NONE NONE = NONE | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 72 | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 73 | fun varify_nonfixed_terms_global nonfixeds tm = tm | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 74 | |> Same.commit (Term_Subst.map_aterms_same | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 75 | (fn Free (x, T) => if member (op =) nonfixeds x then Var ((x, 0), T) else raise Same.SAME | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 76 | | Var (xi, _) => raise TERM (Logic.bad_schematic xi, [tm]) | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 77 | | _ => raise Same.SAME)) | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 78 | |
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 79 | fun max_outcome outcomes = | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 80 | let | 
| 75046 | 81 | val some = find_first (fn (SH_Some _, _) => true | _ => false) outcomes | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 82 | val unknown = find_first (fn (SH_Unknown, _) => true | _ => false) outcomes | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 83 | val timeout = find_first (fn (SH_Timeout, _) => true | _ => false) outcomes | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 84 | val none = find_first (fn (SH_None, _) => true | _ => false) outcomes | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 85 | in | 
| 75046 | 86 | some | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 87 | |> alternative snd unknown | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 88 | |> alternative snd timeout | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 89 | |> alternative snd none | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 90 | |> the_default (SH_Unknown, "") | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 91 | end | 
| 43020 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43006diff
changeset | 92 | |
| 75372 | 93 | fun play_one_line_proofs minimize timeout used_facts state i methss = | 
| 63311 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 94 | (if timeout = Time.zeroTime then | 
| 75372 | 95 | [] | 
| 63311 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 96 | else | 
| 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 97 | let | 
| 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 98 | val ctxt = Proof.context_of state | 
| 75372 | 99 | val used_facts = filter_out (fn (_, (sc, _)) => sc = Chained) used_facts | 
| 100 | val fact_names = map fst used_facts | |
| 63311 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 101 |        val {facts = chained, goal, ...} = Proof.goal state
 | 
| 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 102 | val goal_t = Logic.get_goal (Thm.prop_of goal) i | 
| 57734 
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
 blanchet parents: 
57557diff
changeset | 103 | |
| 75372 | 104 | fun try_methss ress [] = ress | 
| 63311 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 105 | | try_methss ress (meths :: methss) = | 
| 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 106 | let | 
| 75372 | 107 | fun mk_step meths = | 
| 72584 | 108 |                Prove {
 | 
| 109 | qualifiers = [], | |
| 110 | obtains = [], | |
| 111 |                  label = ("", 0),
 | |
| 112 | goal = goal_t, | |
| 113 | subproofs = [], | |
| 114 | facts = ([], fact_names), | |
| 115 | proof_methods = meths, | |
| 116 | comment = ""} | |
| 75372 | 117 | val ress' = | 
| 118 | preplay_isar_step ctxt chained timeout [] (mk_step meths) | |
| 75868 
e7b04452eef3
revived 'try0' and 'smart' Isar proofs in Sledgehammer
 blanchet parents: 
75664diff
changeset | 119 | |> map (fn (meth, play_outcome) => | 
| 75372 | 120 | (case (minimize, play_outcome) of | 
| 121 | (true, Played time) => | |
| 122 | let | |
| 123 | val (time', used_names') = | |
| 124 | minimized_isar_step ctxt chained time (mk_step [meth]) | |
| 125 | ||> (facts_of_isar_step #> snd) | |
| 126 | val used_facts' = filter (member (op =) used_names' o fst) used_facts | |
| 127 | in | |
| 128 | (meth, Played time', used_facts') | |
| 129 | end | |
| 130 | | _ => (meth, play_outcome, used_facts))) | |
| 131 | val any_succeeded = exists (fn (_, Played _, _) => true | _ => false) ress' | |
| 63311 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 132 | in | 
| 75372 | 133 | try_methss (ress' @ ress) (if any_succeeded then [] else methss) | 
| 63311 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 134 | end | 
| 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 135 | in | 
| 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 136 | try_methss [] methss | 
| 
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
 blanchet parents: 
63097diff
changeset | 137 | end) | 
| 76524 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 138 | |> map (fn (meth, play_outcome, used_facts) => | 
| 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 139 | (meth, (play_outcome, filter_out (fn (_, (sc, _)) => sc = Chained) used_facts))) | 
| 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 140 | |> sort (play_outcome_ord o apply2 (fn (_, (play_outcome, _)) => play_outcome)) | 
| 75372 | 141 | |
| 142 | fun select_one_line_proof used_facts preferred_meth preplay_results = | |
| 143 | (case preplay_results of | |
| 75374 
6e8ca4959334
tuned sledehammer to return best succeeding preplay method
 desharna parents: 
75373diff
changeset | 144 | (* Select best method if preplay succeeded *) | 
| 76524 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 145 | (best_meth, (best_outcome as Played _, best_used_facts)) :: _ => | 
| 75374 
6e8ca4959334
tuned sledehammer to return best succeeding preplay method
 desharna parents: 
75373diff
changeset | 146 | (best_used_facts, (best_meth, best_outcome)) | 
| 75868 
e7b04452eef3
revived 'try0' and 'smart' Isar proofs in Sledgehammer
 blanchet parents: 
75664diff
changeset | 147 | (* Otherwise select preferred method *) | 
| 76524 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 148 | | _ => | 
| 75868 
e7b04452eef3
revived 'try0' and 'smart' Isar proofs in Sledgehammer
 blanchet parents: 
75664diff
changeset | 149 | (used_facts, (preferred_meth, | 
| 76524 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 150 | (case AList.lookup (op =) preplay_results preferred_meth of | 
| 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 151 | SOME (outcome, _) => outcome | 
| 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 152 | | NONE => Play_Timed_Out Time.zeroTime)))) | 
| 75372 | 153 | |> apfst (filter_out (fn (_, (sc, _)) => sc = Chained)) | 
| 57734 
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
 blanchet parents: 
57557diff
changeset | 154 | |
| 75035 | 155 | fun launch_prover (params as {verbose, spy, slices, timeout, ...}) mode learn
 | 
| 75033 | 156 |     (problem as {state, subgoal, factss, ...} : prover_problem)
 | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 157 | (slice as ((slice_size, abduce, check_consistent, num_facts, fact_filter), _)) name = | 
| 41089 | 158 | let | 
| 159 | val ctxt = Proof.context_of state | |
| 53800 | 160 | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 161 | val _ = spying spy (fn () => (state, subgoal, name, | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 162 | "Launched" ^ (if abduce then " (abduce)" else "") ^ | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 163 | (if check_consistent then " (check_consistent)" else ""))) | 
| 53800 | 164 | |
| 75033 | 165 | val _ = | 
| 166 | if verbose then | |
| 167 | writeln (name ^ " with " ^ string_of_int num_facts ^ " " ^ fact_filter ^ " fact" ^ | |
| 75340 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 168 | plural_s num_facts ^ " for " ^ string_of_time (slice_timeout slice_size slices timeout) ^ | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 169 | (if abduce then " (abduce)" else "") ^ | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 170 | (if check_consistent then " (check_consistent)" else "") ^ "...") | 
| 75033 | 171 | else | 
| 172 | () | |
| 173 | ||
| 51009 
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
 blanchet parents: 
51008diff
changeset | 174 | fun print_used_facts used_facts used_from = | 
| 
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
 blanchet parents: 
51008diff
changeset | 175 | tag_list 1 used_from | 
| 51005 
ce4290c33d73
eliminated needless speed optimization -- and simplified code quite a bit
 blanchet parents: 
51004diff
changeset | 176 | |> map (fn (j, fact) => fact |> apsnd (K j)) | 
| 48798 | 177 | |> filter_used_facts false used_facts | 
| 48394 
82fc8c956cdc
fixed various issues with MaSh's file handling + tune output + generate local facts again + handle nameless facts gracefully
 blanchet parents: 
48384diff
changeset | 178 | |> map (fn ((name, _), j) => name ^ "@" ^ string_of_int j) | 
| 
82fc8c956cdc
fixed various issues with MaSh's file handling + tune output + generate local facts again + handle nameless facts gracefully
 blanchet parents: 
48384diff
changeset | 179 | |> commas | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 180 |       |> prefix ("Facts in " ^ name ^ " " ^
 | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 181 | (if check_consistent then "inconsistency" else "proof") ^ ": ") | 
| 58843 | 182 | |> writeln | 
| 53800 | 183 | |
| 54062 | 184 |     fun spying_str_of_res ({outcome = NONE, used_facts, used_from, ...} : prover_result) =
 | 
| 185 | let | |
| 186 | val num_used_facts = length used_facts | |
| 54063 | 187 | |
| 188 | fun find_indices facts = | |
| 189 | tag_list 1 facts | |
| 54062 | 190 | |> map (fn (j, fact) => fact |> apsnd (K j)) | 
| 191 | |> filter_used_facts false used_facts | |
| 54773 | 192 | |> distinct (eq_fst (op =)) | 
| 54062 | 193 | |> map (prefix "@" o string_of_int o snd) | 
| 54063 | 194 | |
| 195 | fun filter_info (fact_filter, facts) = | |
| 196 | let | |
| 197 | val indices = find_indices facts | |
| 75033 | 198 | (* "Int.max" is there for robustness *) | 
| 54773 | 199 | val unknowns = replicate (Int.max (0, num_used_facts - length indices)) "?" | 
| 200 | in | |
| 201 | (commas (indices @ unknowns), fact_filter) | |
| 202 | end | |
| 54063 | 203 | |
| 204 | val filter_infos = | |
| 75025 | 205 |             map filter_info (("actual", used_from) :: factss)
 | 
| 54063 | 206 | |> AList.group (op =) | 
| 207 | |> map (fn (indices, fact_filters) => commas fact_filters ^ ": " ^ indices) | |
| 54062 | 208 | in | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 209 | "Success: Found " ^ (if check_consistent then "inconsistency" else "proof") ^ " with " ^ | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 210 | string_of_int num_used_facts ^ " fact" ^ plural_s num_used_facts ^ | 
| 54063 | 211 | (if num_used_facts = 0 then "" else ": " ^ commas filter_infos) | 
| 53800 | 212 | end | 
| 213 |       | spying_str_of_res {outcome = SOME failure, ...} =
 | |
| 54062 | 214 | "Failure: " ^ string_of_atp_failure failure | 
| 74952 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 215 | in | 
| 75033 | 216 | get_minimizing_prover ctxt mode learn name params problem slice | 
| 217 |    |> verbose ? tap (fn {outcome = NONE, used_facts as _ :: _, used_from, ...} =>
 | |
| 218 | print_used_facts used_facts used_from | |
| 219 | | _ => ()) | |
| 220 | |> spy ? tap (fn res => spying spy (fn () => (state, subgoal, name, spying_str_of_res res))) | |
| 74952 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 221 | end | 
| 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 222 | |
| 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 223 | fun preplay_prover_result ({ minimize, preplay_timeout, ...} : params) state subgoal
 | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 224 |     (result as {outcome, used_facts, preferred_methss, message, ...} : prover_result) =
 | 
| 74952 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 225 | let | 
| 75372 | 226 | val (output, chosen_preplay_outcome) = | 
| 74952 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 227 | if outcome = SOME ATP_Proof.TimedOut then | 
| 75372 | 228 | (SH_Timeout, select_one_line_proof used_facts (fst preferred_methss) []) | 
| 74952 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 229 | else if is_some outcome then | 
| 75372 | 230 | (SH_None, select_one_line_proof used_facts (fst preferred_methss) []) | 
| 74952 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 231 | else | 
| 75372 | 232 | let | 
| 233 | val preplay_results = | |
| 234 | play_one_line_proofs minimize preplay_timeout used_facts state subgoal | |
| 235 | (snd preferred_methss) | |
| 236 | val chosen_preplay_outcome = | |
| 237 | select_one_line_proof used_facts (fst preferred_methss) preplay_results | |
| 238 | in | |
| 239 | (SH_Some (result, preplay_results), chosen_preplay_outcome) | |
| 240 | end | |
| 241 | fun output_message () = message (fn () => chosen_preplay_outcome) | |
| 74952 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 242 | in | 
| 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 243 | (output, output_message) | 
| 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 244 | end | 
| 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 245 | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 246 | fun analyze_prover_result_for_consistency (result as {outcome, used_facts, ...} : prover_result) =
 | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 247 | if outcome = SOME ATP_Proof.TimedOut then | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 248 | (SH_Timeout, K "") | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 249 | else if is_some outcome then | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 250 | (SH_None, K "") | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 251 | else | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 252 | (SH_Some (result, []), fn () => | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 253 | (if member (op = o apsnd fst) used_facts sledgehammer_goal_as_fact then | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 254 | (case map fst (filter_out (equal sledgehammer_goal_as_fact o fst) used_facts) of | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 255 | [] => "The goal is inconsistent" | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 256 | | facts => "The goal is inconsistent with these facts: " ^ commas facts) | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 257 | else | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 258 | "The following facts are inconsistent: " ^ | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 259 | commas (map fst used_facts))) | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 260 | |
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 261 | fun check_expected_outcome ctxt prover_name expect outcome = | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 262 | let | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 263 | val outcome_code = short_string_of_sledgehammer_outcome outcome | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 264 | in | 
| 75027 
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
 blanchet parents: 
75026diff
changeset | 265 | (* The "expect" argument is deliberately ignored if the prover is missing so that | 
| 
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
 blanchet parents: 
75026diff
changeset | 266 | "Metis_Examples" can be processed on any machine. *) | 
| 75373 
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
 desharna parents: 
75372diff
changeset | 267 | if expect = "" orelse not (is_prover_installed ctxt prover_name) then | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 268 | () | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 269 | else | 
| 75373 
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
 desharna parents: 
75372diff
changeset | 270 | (case (expect, outcome) of | 
| 
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
 desharna parents: 
75372diff
changeset | 271 |         ("some", SH_Some _) => ()
 | 
| 
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
 desharna parents: 
75372diff
changeset | 272 |       | ("some_preplayed", SH_Some (_, preplay_results)) =>
 | 
| 76524 
87217c655984
correctly show '(> 2 s, timed out)' or similar in Sledgehammer's output
 blanchet parents: 
75874diff
changeset | 273 | if exists (fn (_, (Played _, _)) => true | _ => false) preplay_results then | 
| 75373 
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
 desharna parents: 
75372diff
changeset | 274 | () | 
| 
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
 desharna parents: 
75372diff
changeset | 275 | else | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 276 |             error ("Unexpected outcome: the external prover found a proof but preplay failed")
 | 
| 75373 
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
 desharna parents: 
75372diff
changeset | 277 |       | ("unknown", SH_Unknown) => ()
 | 
| 
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
 desharna parents: 
75372diff
changeset | 278 |       | ("timeout", SH_Timeout) => ()
 | 
| 
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
 desharna parents: 
75372diff
changeset | 279 |       | ("none", SH_None) => ()
 | 
| 
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
 desharna parents: 
75372diff
changeset | 280 |       | _ => error ("Unexpected outcome: " ^ quote outcome_code))
 | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 281 | end | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 282 | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 283 | fun launch_prover_and_preplay (params as {debug, timeout, expect, ...}) mode found_something
 | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 284 |     writeln_result learn (problem as {state, subgoal, ...})
 | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 285 | (slice as ((_, _, check_consistent, _, _), _)) prover_name = | 
| 74952 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 286 | let | 
| 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 287 | val ctxt = Proof.context_of state | 
| 
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
 desharna parents: 
74951diff
changeset | 288 | val hard_timeout = Time.scale 5.0 timeout | 
| 53800 | 289 | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 290 |     fun flip_problem {comment, state, goal, subgoal, factss = factss, ...} =
 | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 291 | let | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 292 | val thy = Proof_Context.theory_of ctxt | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 293 | val assms = Assumption.all_assms_of ctxt | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 294 | val assm_ts = map Thm.term_of assms | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 295 | val subgoal_t = Logic.get_goal (Thm.prop_of goal) subgoal | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 296 | val polymorphic_subgoal_t = (Logic.list_implies (assm_ts, subgoal_t)) | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 297 | |> Logic.varify_global | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 298 | val nonfixeds = | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 299 | subtract (op =) (fold Term.add_free_names assm_ts []) (Term.add_free_names subgoal_t []) | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 300 | val monomorphic_subgoal_t = subgoal_t | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 301 | |> varify_nonfixed_terms_global nonfixeds | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 302 | val subgoal_thms = map (Skip_Proof.make_thm thy) | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 303 | [polymorphic_subgoal_t, monomorphic_subgoal_t] | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 304 | val new_facts = | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 305 | map (fn thm => (((sledgehammer_goal_as_fact, (Assum, General)), thm))) subgoal_thms | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 306 | in | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 307 |         {comment = comment, state = state, goal = Thm.trivial @{cprop False}, subgoal = 1,
 | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 308 | subgoal_count = 1, factss = map (apsnd (append new_facts)) factss, | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 309 | found_something = found_something "an inconsistency"} | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 310 | end | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 311 | |
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 312 | val problem = problem |> check_consistent ? flip_problem | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 313 | |
| 41255 
a80024d7b71b
added debugging option to find out how good the relevance filter was at identifying relevant facts
 blanchet parents: 
41245diff
changeset | 314 | fun really_go () = | 
| 75027 
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
 blanchet parents: 
75026diff
changeset | 315 | launch_prover params mode learn problem slice prover_name | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 316 | |> (if check_consistent then analyze_prover_result_for_consistency else | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 317 | preplay_prover_result params state subgoal) | 
| 53800 | 318 | |
| 41089 | 319 | fun go () = | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 320 | if debug then | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 321 | really_go () | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 322 | else | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 323 | (really_go () | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 324 | handle | 
| 75664 | 325 | ERROR msg => (SH_Unknown, fn () => "Warning: " ^ msg ^ "\n") | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 326 | | exn => | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 327 | if Exn.is_interrupt exn then Exn.reraise exn | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 328 | else (SH_Unknown, fn () => "Internal error:\n" ^ Runtime.exn_message exn ^ "\n")) | 
| 57056 | 329 | |
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 330 | val (outcome, message) = Timeout.apply hard_timeout go () | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 331 | val () = check_expected_outcome ctxt prover_name expect outcome | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 332 | |
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 333 | val message = message () | 
| 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 334 | val () = | 
| 75040 | 335 | if mode = Auto_Try then | 
| 336 | () | |
| 337 | else | |
| 338 | (case outcome of | |
| 339 | SH_Some _ => the_default writeln writeln_result (prover_name ^ ": " ^ message) | |
| 340 | | _ => ()) | |
| 41089 | 341 | in | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 342 | (outcome, message) | 
| 41089 | 343 | end | 
| 344 | ||
| 75034 | 345 | fun string_of_facts filter facts = | 
| 346 | "Selected " ^ string_of_int (length facts) ^ " " ^ (if filter = "" then "" else filter ^ " ") ^ | |
| 347 | "fact" ^ plural_s (length facts) ^ ": " ^ (space_implode " " (map (fst o fst) facts)) | |
| 51008 | 348 | |
| 51010 | 349 | fun string_of_factss factss = | 
| 57464 | 350 | if forall (null o snd) factss then | 
| 63692 | 351 | "Found no relevant facts" | 
| 57464 | 352 | else | 
| 75034 | 353 | cat_lines (map (fn (filter, facts) => string_of_facts filter facts) factss) | 
| 51008 | 354 | |
| 75029 | 355 | val default_slice_schedule = | 
| 75872 
8bfad7bc74cb
tweak Sledgehammer's slicing mechanism -- updated Zipperposition's slices and make them half as long as other provers' to pack more of them in 30 s
 blanchet parents: 
75868diff
changeset | 356 | (* FUDGE (loosely inspired by Seventeen evaluation) *) | 
| 75874 | 357 | [cvc4N, zipperpositionN, vampireN, veritN, spassN, zipperpositionN, eN, cvc4N, | 
| 75872 
8bfad7bc74cb
tweak Sledgehammer's slicing mechanism -- updated Zipperposition's slices and make them half as long as other provers' to pack more of them in 30 s
 blanchet parents: 
75868diff
changeset | 358 | zipperpositionN, cvc4N, zipperpositionN, vampireN, cvc4N, cvc4N, vampireN, | 
| 75874 | 359 | cvc4N, iproverN, zipperpositionN, spassN, vampireN, zipperpositionN, | 
| 360 | vampireN, zipperpositionN, z3N, zipperpositionN, vampireN, iproverN, spassN, | |
| 361 | zipperpositionN, vampireN, cvc4N, zipperpositionN, z3N, z3N, cvc4N, cvc4N, | |
| 362 | zipperpositionN] | |
| 75029 | 363 | |
| 364 | fun schedule_of_provers provers num_slices = | |
| 365 | let | |
| 75032 
8d08bc7e8f98
tweak padding of prover slice schedule to include all provers
 blanchet parents: 
75031diff
changeset | 366 | val (known_provers, unknown_provers) = | 
| 
8d08bc7e8f98
tweak padding of prover slice schedule to include all provers
 blanchet parents: 
75031diff
changeset | 367 | List.partition (member (op =) default_slice_schedule) provers | 
| 75029 | 368 | |
| 75034 | 369 | val default_slice_schedule = filter (member (op =) known_provers) default_slice_schedule | 
| 370 | val num_default_slices = length default_slice_schedule | |
| 371 | ||
| 75029 | 372 | fun round_robin _ [] = [] | 
| 373 | | round_robin 0 _ = [] | |
| 374 | | round_robin n (prover :: provers) = prover :: round_robin (n - 1) (provers @ [prover]) | |
| 375 | in | |
| 376 | if num_slices <= num_default_slices then | |
| 377 | take num_slices default_slice_schedule | |
| 378 | else | |
| 75032 
8d08bc7e8f98
tweak padding of prover slice schedule to include all provers
 blanchet parents: 
75031diff
changeset | 379 | default_slice_schedule | 
| 
8d08bc7e8f98
tweak padding of prover slice schedule to include all provers
 blanchet parents: 
75031diff
changeset | 380 | @ round_robin (num_slices - num_default_slices) (unknown_provers @ known_provers) | 
| 75029 | 381 | end | 
| 382 | ||
| 75060 
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
 blanchet parents: 
75056diff
changeset | 383 | fun prover_slices_of_schedule ctxt factss | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 384 |     ({abduce, check_consistent, max_facts, fact_filter, type_enc, lam_trans, uncurried_aliases,
 | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 385 | ...} : params) | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 386 | schedule = | 
| 75029 | 387 | let | 
| 388 | fun triplicate_slices original = | |
| 389 | let | |
| 390 | val shift = | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 391 | map (apfst (fn (slice_size, abduce, check_consistent, num_facts, fact_filter) => | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 392 | (slice_size, abduce, check_consistent, num_facts, | 
| 75339 
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
 blanchet parents: 
75076diff
changeset | 393 | if fact_filter = mashN then mepoN | 
| 
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
 blanchet parents: 
75076diff
changeset | 394 | else if fact_filter = mepoN then meshN | 
| 
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
 blanchet parents: 
75076diff
changeset | 395 | else mashN))) | 
| 75029 | 396 | |
| 397 | val shifted_once = shift original | |
| 398 | val shifted_twice = shift shifted_once | |
| 399 | in | |
| 400 | original @ shifted_once @ shifted_twice | |
| 401 | end | |
| 402 | ||
| 75063 
7ff39293e265
added possibility of extra options to SMT slices
 blanchet parents: 
75060diff
changeset | 403 | fun adjust_extra (ATP_Slice (format0, type_enc0, lam_trans0, uncurried_aliases0, | 
| 
7ff39293e265
added possibility of extra options to SMT slices
 blanchet parents: 
75060diff
changeset | 404 | extra_extra0)) = | 
| 
7ff39293e265
added possibility of extra options to SMT slices
 blanchet parents: 
75060diff
changeset | 405 | ATP_Slice (format0, the_default type_enc0 type_enc, the_default lam_trans0 lam_trans, | 
| 
7ff39293e265
added possibility of extra options to SMT slices
 blanchet parents: 
75060diff
changeset | 406 | the_default uncurried_aliases0 uncurried_aliases, extra_extra0) | 
| 
7ff39293e265
added possibility of extra options to SMT slices
 blanchet parents: 
75060diff
changeset | 407 | | adjust_extra (extra as SMT_Slice _) = extra | 
| 75034 | 408 | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 409 | fun adjust_slice max_slice_size | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 410 | ((slice_size0, abduce0, check_consistent0, num_facts0, fact_filter0), extra) = | 
| 75060 
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
 blanchet parents: 
75056diff
changeset | 411 | let | 
| 75340 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 412 | val slice_size = Int.min (max_slice_size, slice_size0) | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 413 | val abduce = abduce |> the_default abduce0 | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 414 | val check_consistent = check_consistent |> the_default check_consistent0 | 
| 75060 
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
 blanchet parents: 
75056diff
changeset | 415 | val fact_filter = fact_filter |> the_default fact_filter0 | 
| 
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
 blanchet parents: 
75056diff
changeset | 416 | val max_facts = max_facts |> the_default num_facts0 | 
| 
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
 blanchet parents: 
75056diff
changeset | 417 | val num_facts = Int.min (max_facts, length (facts_of_filter fact_filter factss)) | 
| 
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
 blanchet parents: 
75056diff
changeset | 418 | in | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 419 | ((slice_size, abduce, check_consistent, num_facts, fact_filter), adjust_extra extra) | 
| 75060 
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
 blanchet parents: 
75056diff
changeset | 420 | end | 
| 75034 | 421 | |
| 75029 | 422 | val provers = distinct (op =) schedule | 
| 423 | val prover_slices = | |
| 75034 | 424 | map (fn prover => (prover, | 
| 75340 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 425 | (is_none fact_filter ? triplicate_slices) (get_slices ctxt prover))) | 
| 75034 | 426 | provers | 
| 75029 | 427 | |
| 75340 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 428 | val max_threads = Multithreading.max_threads () | 
| 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 429 | |
| 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 430 | fun translate_schedule _ 0 _ = [] | 
| 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 431 | | translate_schedule _ _ [] = [] | 
| 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 432 | | translate_schedule prover_slices slices_left (prover :: schedule) = | 
| 75029 | 433 | (case AList.lookup (op =) prover_slices prover of | 
| 75340 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 434 | SOME (slice0 :: slices) => | 
| 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 435 | let | 
| 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 436 | val prover_slices' = AList.update (op =) (prover, slices) prover_slices | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 437 | val slice as ((slice_size, _, _, _, _), _) = | 
| 75340 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 438 | adjust_slice ((slices_left + max_threads - 1) div max_threads) slice0 | 
| 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 439 | in | 
| 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 440 | (prover, slice) :: translate_schedule prover_slices' (slices_left - slice_size) schedule | 
| 75029 | 441 | end | 
| 75340 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 442 | | _ => translate_schedule prover_slices slices_left schedule) | 
| 75029 | 443 | in | 
| 75340 
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
 blanchet parents: 
75339diff
changeset | 444 | translate_schedule prover_slices (length schedule) schedule | 
| 75034 | 445 | |> distinct (op =) | 
| 75029 | 446 | end | 
| 447 | ||
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 448 | fun run_sledgehammer (params as {verbose, spy, provers, check_consistent, induction_rules,
 | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 449 | max_facts, max_proofs, slices, ...}) | 
| 75030 | 450 |     mode writeln_result i (fact_override as {only, ...}) state =
 | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39494diff
changeset | 451 | if null provers then | 
| 63692 | 452 | error "No prover is set" | 
| 55286 | 453 | else | 
| 454 | (case subgoal_count state of | |
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 455 | 0 => (error "No subgoal!"; (false, (SH_None, ""))) | 
| 55286 | 456 | | n => | 
| 457 | let | |
| 458 | val _ = Proof.assert_backward state | |
| 62735 | 459 | val print = if mode = Normal andalso is_none writeln_result then writeln else K () | 
| 460 | ||
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 461 | val found_proofs_and_inconsistencies = Synchronized.var "found_proofs_and_inconsistencies" 0 | 
| 75030 | 462 | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 463 | fun found_something a_proof_or_inconsistency prover_name = | 
| 75030 | 464 | if mode = Normal then | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 465 | (Synchronized.change found_proofs_and_inconsistencies (fn n => n + 1); | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 466 | (the_default writeln writeln_result) (prover_name ^ " found " ^ | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 467 | a_proof_or_inconsistency ^ "...")) | 
| 75030 | 468 | else | 
| 469 | () | |
| 62735 | 470 | |
| 55286 | 471 | val ctxt = Proof.context_of state | 
| 74950 
b350a1f2115d
added nearly_all_facts_of_context and uniformized its usage in Sledgehammer and Mirabelle
 desharna parents: 
73975diff
changeset | 472 | val inst_inducts = induction_rules = SOME Instantiate | 
| 
b350a1f2115d
added nearly_all_facts_of_context and uniformized its usage in Sledgehammer and Mirabelle
 desharna parents: 
73975diff
changeset | 473 |         val {facts = chained_thms, goal, ...} = Proof.goal state
 | 
| 55286 | 474 | val (_, hyp_ts, concl_t) = strip_subgoal goal i ctxt | 
| 475 | val _ = | |
| 476 | (case find_first (not o is_prover_supported ctxt) provers of | |
| 63692 | 477 |             SOME name => error ("No such prover: " ^ name)
 | 
| 55286 | 478 | | NONE => ()) | 
| 479 | val _ = print "Sledgehammering..." | |
| 57037 | 480 | val _ = spying spy (fn () => (state, i, "***", "Starting " ^ str_of_mode mode ^ " mode")) | 
| 75002 | 481 |         val ({elapsed, ...}, all_facts) = Timing.timing
 | 
| 482 | (nearly_all_facts_of_context ctxt inst_inducts fact_override chained_thms hyp_ts) concl_t | |
| 483 | val _ = spying spy (fn () => (state, i, "All", | |
| 484 | "Extracting " ^ string_of_int (length all_facts) ^ " facts from background theory in " ^ | |
| 485 | string_of_int (Time.toMilliseconds elapsed) ^ " ms")) | |
| 54090 
a28992e35032
run relevance filter only once for ATPs and SMT solvers, since it should now yield the same results anyway
 blanchet parents: 
54063diff
changeset | 486 | |
| 55286 | 487 | val spying_str_of_factss = | 
| 488 | commas o map (fn (filter, facts) => filter ^ ": " ^ string_of_int (length facts)) | |
| 53800 | 489 | |
| 55286 | 490 | fun get_factss provers = | 
| 491 | let | |
| 492 | val max_max_facts = | |
| 493 | (case max_facts of | |
| 494 | SOME n => n | |
| 495 | | NONE => | |
| 75029 | 496 | fold (fn prover => | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 497 | fold (fn ((_, _, _, max_facts, _), _) => Integer.max max_facts) | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 498 | (get_slices ctxt prover)) | 
| 75029 | 499 | provers 0) | 
| 75034 | 500 | * 51 div 50 (* some slack to account for filtering of induction facts below *) | 
| 75027 
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
 blanchet parents: 
75026diff
changeset | 501 | |
| 74998 
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
 desharna parents: 
74953diff
changeset | 502 |             val ({elapsed, ...}, factss) = Timing.timing
 | 
| 
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
 desharna parents: 
74953diff
changeset | 503 | (relevant_facts ctxt params (hd provers) max_max_facts fact_override hyp_ts concl_t) | 
| 
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
 desharna parents: 
74953diff
changeset | 504 | all_facts | 
| 75027 
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
 blanchet parents: 
75026diff
changeset | 505 | |
| 
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
 blanchet parents: 
75026diff
changeset | 506 | val induction_rules = the_default (if only then Include else Exclude) induction_rules | 
| 
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
 blanchet parents: 
75026diff
changeset | 507 | val factss = map (apsnd (maybe_filter_out_induction_rules induction_rules)) factss | 
| 
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
 blanchet parents: 
75026diff
changeset | 508 | |
| 74998 
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
 desharna parents: 
74953diff
changeset | 509 | val () = spying spy (fn () => (state, i, "All", | 
| 75002 | 510 | "Filtering facts in " ^ string_of_int (Time.toMilliseconds elapsed) ^ | 
| 511 | " ms (MaSh algorithm: " ^ str_of_mash_algorithm (the_mash_algorithm ()) ^ ")")); | |
| 74998 
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
 desharna parents: 
74953diff
changeset | 512 | val () = if verbose then print (string_of_factss factss) else () | 
| 
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
 desharna parents: 
74953diff
changeset | 513 | val () = spying spy (fn () => | 
| 
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
 desharna parents: 
74953diff
changeset | 514 | (state, i, "All", "Selected facts: " ^ spying_str_of_factss factss)) | 
| 55286 | 515 | in | 
| 74998 
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
 desharna parents: 
74953diff
changeset | 516 | factss | 
| 55286 | 517 | end | 
| 53800 | 518 | |
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58843diff
changeset | 519 | fun launch_provers () = | 
| 55286 | 520 | let | 
| 75060 
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
 blanchet parents: 
75056diff
changeset | 521 | val factss = get_factss provers | 
| 55286 | 522 | val problem = | 
| 523 |               {comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
 | |
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 524 | factss = factss, found_something = found_something "a proof"} | 
| 69706 
6d6235b828fc
get rid of visibility in MaSh -- it slows it down more than it helps
 blanchet parents: 
63697diff
changeset | 525 | val learn = mash_learn_proof ctxt params (Thm.prop_of goal) | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 526 | val launch = launch_prover_and_preplay params mode found_something writeln_result learn | 
| 75029 | 527 | |
| 528 | val schedule = | |
| 75036 | 529 | if mode = Auto_Try then provers | 
| 530 | else schedule_of_provers provers slices | |
| 75060 
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
 blanchet parents: 
75056diff
changeset | 531 | val prover_slices = prover_slices_of_schedule ctxt factss params schedule | 
| 75037 | 532 | |
| 533 | val _ = | |
| 534 | if verbose then | |
| 535 |                 writeln ("Running " ^ commas (map fst prover_slices) ^ "...")
 | |
| 536 | else | |
| 537 | () | |
| 55286 | 538 | in | 
| 539 | if mode = Auto_Try then | |
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 540 | (SH_Unknown, "") | 
| 75029 | 541 | |> fold (fn (prover, slice) => | 
| 75025 | 542 | fn accum as (SH_Some _, _) => accum | 
| 75029 | 543 | | _ => launch problem slice prover) | 
| 544 | prover_slices | |
| 55286 | 545 | else | 
| 74950 
b350a1f2115d
added nearly_all_facts_of_context and uniformized its usage in Sledgehammer and Mirabelle
 desharna parents: 
73975diff
changeset | 546 | (learn chained_thms; | 
| 75031 | 547 | Par_List.map (fn (prover, slice) => | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 548 | if Synchronized.value found_proofs_and_inconsistencies < max_proofs then | 
| 75031 | 549 | launch problem slice prover | 
| 550 | else | |
| 75056 
04a4881ff0fd
propagate right result when enough proofs have been found
 blanchet parents: 
75054diff
changeset | 551 | (SH_None, "")) | 
| 75031 | 552 | prover_slices | 
| 74953 
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
 desharna parents: 
74952diff
changeset | 553 | |> max_outcome) | 
| 55286 | 554 | end | 
| 555 | in | |
| 75020 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 556 | (launch_provers () | 
| 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 557 | handle Timeout.TIMEOUT _ => (SH_Timeout, "")) | 
| 75054 | 558 | |> `(fn (outcome, message) => | 
| 75020 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 559 | (case outcome of | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 560 | SH_Some _ => (the_default writeln writeln_result "Done"; true) | 
| 75076 | 561 | | SH_Unknown => (the_default writeln writeln_result message; false) | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 562 | | SH_Timeout => | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 563 | (the_default writeln writeln_result | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 564 |                ("No " ^ (if check_consistent = SOME true then "inconsistency" else "proof") ^
 | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 565 | " found"); | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 566 | false) | 
| 75076 | 567 | | SH_None => (the_default writeln writeln_result | 
| 77269 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 568 | (if message = "" then | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 569 | "No " ^ (if check_consistent = SOME true then "inconsistency" else "proof") ^ | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 570 | " found" | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 571 | else | 
| 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 blanchet parents: 
76524diff
changeset | 572 | "Warning: " ^ message); | 
| 75075 | 573 | false))) | 
| 75020 
b087610592b4
rationalized output for forthcoming slicing model
 blanchet parents: 
75019diff
changeset | 574 | end) | 
| 38044 | 575 | |
| 28582 | 576 | end; |