author | Fabian Huch <huch@in.tum.de> |
Tue, 07 Jun 2022 17:20:25 +0200 | |
changeset 75521 | 7a289e681454 |
parent 75341 | 72cbbb4d98f3 |
child 75868 | e7b04452eef3 |
permissions | -rw-r--r-- |
55205 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML |
2 |
Author: Fabian Immler, TU Muenchen |
|
3 |
Author: Makarius |
|
4 |
Author: Jasmin Blanchette, TU Muenchen |
|
5 |
||
6 |
ATPs as Sledgehammer provers. |
|
7 |
*) |
|
8 |
||
9 |
signature SLEDGEHAMMER_PROVER_ATP = |
|
10 |
sig |
|
11 |
type mode = Sledgehammer_Prover.mode |
|
12 |
type prover = Sledgehammer_Prover.prover |
|
13 |
||
74981
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
14 |
val atp_problem_dest_dir : string Config.T |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
15 |
val atp_proof_dest_dir : string Config.T |
55212 | 16 |
val atp_problem_prefix : string Config.T |
62718 | 17 |
val atp_completish : int Config.T |
55212 | 18 |
val atp_full_names : bool Config.T |
19 |
||
55205 | 20 |
val run_atp : mode -> string -> prover |
21 |
end; |
|
22 |
||
23 |
structure Sledgehammer_Prover_ATP : SLEDGEHAMMER_PROVER_ATP = |
|
24 |
struct |
|
25 |
||
26 |
open ATP_Util |
|
27 |
open ATP_Problem |
|
72399
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
desharna
parents:
71931
diff
changeset
|
28 |
open ATP_Problem_Generate |
55205 | 29 |
open ATP_Proof |
30 |
open ATP_Proof_Reconstruct |
|
31 |
open Sledgehammer_Util |
|
55287 | 32 |
open Sledgehammer_Proof_Methods |
55205 | 33 |
open Sledgehammer_Isar |
72400 | 34 |
open Sledgehammer_ATP_Systems |
55205 | 35 |
open Sledgehammer_Prover |
36 |
||
55212 | 37 |
(* Empty string means create files in Isabelle's temporary files directory. *) |
74981
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
38 |
val atp_problem_dest_dir = |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
39 |
Attrib.setup_config_string \<^binding>\<open>sledgehammer_atp_problem_dest_dir\<close> (K "") |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
40 |
val atp_proof_dest_dir = |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
41 |
Attrib.setup_config_string \<^binding>\<open>sledgehammer_atp_proof_dest_dir\<close> (K "") |
55212 | 42 |
val atp_problem_prefix = |
69593 | 43 |
Attrib.setup_config_string \<^binding>\<open>sledgehammer_atp_problem_prefix\<close> (K "prob") |
44 |
val atp_completish = Attrib.setup_config_int \<^binding>\<open>sledgehammer_atp_completish\<close> (K 0) |
|
55212 | 45 |
(* In addition to being easier to read, readable names are often much shorter, especially if types |
46 |
are mangled in names. This makes a difference for some provers (e.g., E). For these reason, short |
|
47 |
names are enabled by default. *) |
|
69593 | 48 |
val atp_full_names = Attrib.setup_config_bool \<^binding>\<open>sledgehammer_atp_full_names\<close> (K false) |
55212 | 49 |
|
75035 | 50 |
fun choose_type_enc strictness format good_type_enc = |
51 |
type_enc_of_string strictness good_type_enc |
|
52 |
|> adjust_type_enc format |
|
55205 | 53 |
|
54 |
fun has_bound_or_var_of_type pred = |
|
55 |
exists_subterm (fn Var (_, T as Type _) => pred T |
|
56 |
| Abs (_, T as Type _, _) => pred T |
|
57 |
| _ => false) |
|
58 |
||
59 |
(* Unwanted equalities are those between a (bound or schematic) variable that does not properly |
|
60 |
occur in the second operand. *) |
|
61 |
val is_exhaustive_finite = |
|
62 |
let |
|
63 |
fun is_bad_equal (Var z) t = |
|
64 |
not (exists_subterm (fn Var z' => z = z' | _ => false) t) |
|
65 |
| is_bad_equal (Bound j) t = not (loose_bvar1 (t, j)) |
|
66 |
| is_bad_equal _ _ = false |
|
67 |
fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1 |
|
68 |
fun do_formula pos t = |
|
55208 | 69 |
(case (pos, t) of |
74379 | 70 |
(_, \<^Const_>\<open>Trueprop for t1\<close>) => do_formula pos t1 |
69593 | 71 |
| (true, Const (\<^const_name>\<open>Pure.all\<close>, _) $ Abs (_, _, t')) => do_formula pos t' |
72 |
| (true, Const (\<^const_name>\<open>All\<close>, _) $ Abs (_, _, t')) => do_formula pos t' |
|
73 |
| (false, Const (\<^const_name>\<open>Ex\<close>, _) $ Abs (_, _, t')) => do_formula pos t' |
|
74379 | 74 |
| (_, \<^Const_>\<open>Pure.imp for t1 t2\<close>) => |
69593 | 75 |
do_formula (not pos) t1 andalso (t2 = \<^prop>\<open>False\<close> orelse do_formula pos t2) |
74379 | 76 |
| (_, \<^Const_>\<open>implies for t1 t2\<close>) => |
77 |
do_formula (not pos) t1 andalso (t2 = \<^Const>\<open>False\<close> orelse do_formula pos t2) |
|
78 |
| (_, \<^Const_>\<open>Not for t1\<close>) => do_formula (not pos) t1 |
|
79 |
| (true, \<^Const_>\<open>disj for t1 t2\<close>) => forall (do_formula pos) [t1, t2] |
|
80 |
| (false, \<^Const_>\<open>conj for t1 t2\<close>) => forall (do_formula pos) [t1, t2] |
|
69593 | 81 |
| (true, Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t1 $ t2) => do_equals t1 t2 |
82 |
| (true, Const (\<^const_name>\<open>Pure.eq\<close>, _) $ t1 $ t2) => do_equals t1 t2 |
|
55208 | 83 |
| _ => false) |
55205 | 84 |
in do_formula true end |
85 |
||
86 |
(* Facts containing variables of finite types such as "unit" or "bool" or of the form |
|
87 |
"ALL x. x = A | x = B | x = C" are likely to lead to untypable proofs for unsound type |
|
88 |
encodings. *) |
|
89 |
fun is_dangerous_prop ctxt = |
|
90 |
transform_elim_prop |
|
91 |
#> (has_bound_or_var_of_type (is_type_surely_finite ctxt) orf is_exhaustive_finite) |
|
92 |
||
93 |
val mono_max_privileged_facts = 10 |
|
94 |
||
95 |
fun suffix_of_mode Auto_Try = "_try" |
|
96 |
| suffix_of_mode Try = "_try" |
|
97 |
| suffix_of_mode Normal = "" |
|
98 |
| suffix_of_mode MaSh = "" |
|
99 |
| suffix_of_mode Minimize = "_min" |
|
100 |
||
57738 | 101 |
(* Important messages are important but not so important that users want to see them each time. *) |
55205 | 102 |
val atp_important_message_keep_quotient = 25 |
103 |
||
104 |
fun run_atp mode name |
|
75067
c23aa0d177b4
used max_facts and fact_filter from slice for both ATP and SMT in sledgehammer
desharna
parents:
75064
diff
changeset
|
105 |
({debug, verbose, overlord, strict, max_mono_iters, max_new_mono_instances, isar_proofs, |
c23aa0d177b4
used max_facts and fact_filter from slice for both ATP and SMT in sledgehammer
desharna
parents:
75064
diff
changeset
|
106 |
compress, try0, smt_proofs, minimize, slices, timeout, preplay_timeout, spy, ...} : params) |
75025 | 107 |
({comment, state, goal, subgoal, subgoal_count, factss, found_proof} : prover_problem) |
108 |
slice = |
|
55205 | 109 |
let |
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
110 |
val (basic_slice as (slice_size, _, _), |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
111 |
ATP_Slice (good_format, good_type_enc, good_lam_trans, good_uncurried_aliases, extra)) = |
75069
455d886009b1
uniformized fact selection for ATP and SMT in Sledgehammer
desharna
parents:
75067
diff
changeset
|
112 |
slice |
455d886009b1
uniformized fact selection for ATP and SMT in Sledgehammer
desharna
parents:
75067
diff
changeset
|
113 |
val facts = facts_of_basic_slice basic_slice factss |
55205 | 114 |
val thy = Proof.theory_of state |
115 |
val ctxt = Proof.context_of state |
|
116 |
||
75026 | 117 |
val {exec, arguments, proof_delims, known_failures, prem_role, good_max_mono_iters, |
118 |
good_max_new_mono_instances, ...} = get_atp thy name () |
|
55205 | 119 |
|
57671
dc5e1b1db9ba
avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
blanchet
parents:
57460
diff
changeset
|
120 |
val full_proofs = isar_proofs |> the_default (mode = Minimize) |
57267
8b87114357bd
integrated new Waldmeister code with 'sledgehammer' command
blanchet
parents:
57266
diff
changeset
|
121 |
val local_name = perhaps (try (unprefix remote_prefix)) name |
8b87114357bd
integrated new Waldmeister code with 'sledgehammer' command
blanchet
parents:
57266
diff
changeset
|
122 |
|
62718 | 123 |
val completish = Config.get ctxt atp_completish |
124 |
val atp_mode = if completish > 0 then Sledgehammer_Completish completish else Sledgehammer |
|
55205 | 125 |
val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal ctxt |
74981
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
126 |
val (problem_dest_dir, proof_dest_dir, problem_prefix) = |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
127 |
if overlord then |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
128 |
overlord_file_location_of_prover name |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
129 |
|> (fn (dir, prefix) => (dir, dir, prefix)) |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
130 |
else |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
131 |
(Config.get ctxt atp_problem_dest_dir, |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
132 |
Config.get ctxt atp_proof_dest_dir, |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
133 |
Config.get ctxt atp_problem_prefix) |
55205 | 134 |
val problem_file_name = |
135 |
Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^ |
|
58085 | 136 |
suffix_of_mode mode ^ "_" ^ string_of_int subgoal) |
74893 | 137 |
|> Path.ext "p" |
55205 | 138 |
val prob_path = |
74981
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
139 |
if problem_dest_dir = "" then |
55205 | 140 |
File.tmp_path problem_file_name |
74981
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
141 |
else if File.exists (Path.explode problem_dest_dir) then |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
142 |
Path.explode problem_dest_dir + problem_file_name |
55205 | 143 |
else |
74981
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
144 |
error ("No such directory: " ^ quote problem_dest_dir) |
73435
1cc848548f21
invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents:
73432
diff
changeset
|
145 |
val executable = |
55208 | 146 |
(case find_first (fn var => getenv var <> "") (fst exec) of |
55205 | 147 |
SOME var => |
148 |
let |
|
149 |
val pref = getenv var ^ "/" |
|
60978 | 150 |
val paths = |
151 |
map (Path.explode o prefix pref) |
|
152 |
(if ML_System.platform_is_windows then |
|
153 |
map (suffix ".exe") (snd exec) @ snd exec |
|
154 |
else snd exec); |
|
55205 | 155 |
in |
55208 | 156 |
(case find_first File.exists paths of |
55205 | 157 |
SOME path => path |
63692 | 158 |
| NONE => error ("Bad executable: " ^ Path.print (hd paths))) |
55205 | 159 |
end |
63692 | 160 |
| NONE => error ("The environment variable " ^ quote (List.last (fst exec)) ^ " is not set")) |
55205 | 161 |
|
162 |
fun run () = |
|
163 |
let |
|
164 |
fun monomorphize_facts facts = |
|
165 |
let |
|
166 |
val ctxt = |
|
167 |
ctxt |
|
75026 | 168 |
|> repair_monomorph_context max_mono_iters good_max_mono_iters max_new_mono_instances |
169 |
good_max_new_mono_instances |
|
55205 | 170 |
(* pseudo-theorem involving the same constants as the subgoal *) |
171 |
val subgoal_th = |
|
172 |
Logic.list_implies (hyp_ts, concl_t) |> Skip_Proof.make_thm thy |
|
173 |
val rths = |
|
174 |
facts |> chop mono_max_privileged_facts |
|
175 |
|>> map (pair 1 o snd) |
|
176 |
||> map (pair 2 o snd) |
|
177 |
|> op @ |
|
178 |
|> cons (0, subgoal_th) |
|
179 |
in |
|
180 |
Monomorph.monomorph atp_schematic_consts_of ctxt rths |
|
181 |
|> tl |> curry ListPair.zip (map fst facts) |
|
55208 | 182 |
|> maps (fn (name, rths) => map (pair name o zero_var_indexes o snd) rths) |
55205 | 183 |
end |
184 |
||
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
185 |
val strictness = if strict then Strict else Non_Strict |
75035 | 186 |
val type_enc = choose_type_enc strictness good_format good_type_enc |
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
187 |
val run_timeout = slice_timeout slice_size slices timeout |
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
188 |
val generous_run_timeout = if mode = MaSh then one_day else run_timeout |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
189 |
val ({elapsed, ...}, atp_problem_data as (atp_problem, _, _, _)) = Timing.timing (fn () => |
55205 | 190 |
let |
75026 | 191 |
val generate_info = (case good_format of DFG _ => true | _ => false) |
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
192 |
val readable_names = not (Config.get ctxt atp_full_names) |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
193 |
in |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
194 |
facts |
75069
455d886009b1
uniformized fact selection for ATP and SMT in Sledgehammer
desharna
parents:
75067
diff
changeset
|
195 |
|> not (is_type_enc_sound type_enc) ? |
455d886009b1
uniformized fact selection for ATP and SMT in Sledgehammer
desharna
parents:
75067
diff
changeset
|
196 |
filter_out (is_dangerous_prop ctxt o Thm.prop_of o snd) |
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
197 |
|> not (is_type_enc_polymorphic type_enc) ? monomorphize_facts |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
198 |
|> map (apsnd Thm.prop_of) |
75026 | 199 |
|> generate_atp_problem ctxt generate_info good_format prem_role type_enc atp_mode |
75035 | 200 |
good_lam_trans good_uncurried_aliases readable_names true hyp_ts concl_t |
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
201 |
end) () |
75000 | 202 |
|
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
203 |
val () = spying spy (fn () => |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
204 |
(state, subgoal, name, "Generating ATP problem in " ^ |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
205 |
string_of_int (Time.toMilliseconds elapsed) ^ " ms")) |
55205 | 206 |
|
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
207 |
val args = arguments ctxt full_proofs extra run_timeout prob_path |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
208 |
val command = space_implode " " (File.bash_path executable :: args) |
73435
1cc848548f21
invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents:
73432
diff
changeset
|
209 |
|
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
210 |
fun run_command () = |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
211 |
if exec = isabelle_scala_function then |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
212 |
let val {output, timing} = SystemOnTPTP.run_system_encoded args |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
213 |
in (output, timing) end |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
214 |
else |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
215 |
let val res = Isabelle_System.bash_process (Bash.script command |> Bash.redirect) |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
216 |
in (Process_Result.out res, Process_Result.timing_elapsed res) end |
57263 | 217 |
|
75341 | 218 |
val _ = atp_problem |
219 |
|> lines_of_atp_problem good_format (fn () => atp_problem_term_order_info atp_problem) |
|
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
220 |
|> (exec <> isabelle_scala_function) ? |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
221 |
cons ("% " ^ command ^ "\n" ^ (if comment = "" then "" else "% " ^ comment ^ "\n")) |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
222 |
|> File.write_list prob_path |
75000 | 223 |
|
75064 | 224 |
val ((output, run_time), ((atp_proof, tstplike_proof), outcome)) = |
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
225 |
Timeout.apply generous_run_timeout run_command () |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
226 |
|>> overlord ? |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
227 |
(fn output => prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n") output) |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
228 |
|> (fn accum as (output, _) => |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
229 |
(accum, |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
230 |
extract_tstplike_proof_and_outcome verbose proof_delims known_failures output |
75064 | 231 |
|>> `(atp_proof_of_tstplike_proof false (perhaps (try (unprefix remote_prefix)) name) |
232 |
atp_problem) |
|
233 |
handle UNRECOGNIZED_ATP_PROOF () => (([], ""), SOME ProofUnparsable))) |
|
234 |
handle Timeout.TIMEOUT _ => (("", run_timeout), (([], ""), SOME TimedOut)) |
|
235 |
| ERROR msg => (("", Time.zeroTime), (([], ""), SOME (UnknownError msg))) |
|
55205 | 236 |
|
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
237 |
val () = spying spy (fn () => |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
238 |
(state, subgoal, name, "Running command in " ^ |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
239 |
string_of_int (Time.toMilliseconds run_time) ^ " ms")) |
55205 | 240 |
|
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
241 |
val outcome = |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
242 |
(case outcome of |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
243 |
NONE => |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
244 |
(case used_facts_in_unsound_atp_proof ctxt (map fst facts) atp_proof of |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
245 |
SOME facts => |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
246 |
let |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
247 |
val failure = UnsoundProof (is_type_enc_sound type_enc, sort string_ord facts) |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
248 |
in |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
249 |
if debug then (warning (string_of_atp_failure failure); NONE) else SOME failure |
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
250 |
end |
75020
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
251 |
| NONE => (found_proof name; NONE)) |
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
252 |
| _ => outcome) |
55205 | 253 |
in |
75064 | 254 |
(atp_problem_data, |
255 |
(output, run_time, facts, atp_problem, tstplike_proof, atp_proof, outcome), |
|
256 |
(good_format, type_enc)) |
|
55205 | 257 |
end |
258 |
||
259 |
(* If the problem file has not been exported, remove it; otherwise, export |
|
260 |
the proof file too. *) |
|
74981
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
261 |
fun clean_up () = if problem_dest_dir = "" then (try File.rm prob_path; ()) else () |
75064 | 262 |
fun export (_, (output, _, _, _, _, _, _), _) = |
74893 | 263 |
let |
74981
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
264 |
val proof_dest_dir_path = Path.explode proof_dest_dir |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
265 |
val make_export_file_name = |
74893 | 266 |
Path.split_ext |
267 |
#> apfst (Path.explode o suffix "_proof" o Path.implode) |
|
268 |
#> swap |
|
269 |
#> uncurry Path.ext |
|
270 |
in |
|
74981
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
271 |
if proof_dest_dir = "" then Output.system_message "don't export proof" |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
272 |
else if File.exists proof_dest_dir_path then |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
273 |
File.write (proof_dest_dir_path + make_export_file_name problem_file_name) output |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
274 |
else |
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74927
diff
changeset
|
275 |
error ("No such directory: " ^ quote proof_dest_dir) |
74893 | 276 |
end |
55208 | 277 |
|
75064 | 278 |
val ((_, pool, lifted, sym_tab), |
279 |
(output, run_time, used_from, atp_problem, tstplike_proof, atp_proof, outcome), |
|
75018
fcfd96a59625
disable slicing within ATP module (in preparation for refactoring)
blanchet
parents:
75017
diff
changeset
|
280 |
(format, type_enc)) = |
55205 | 281 |
with_cleanup clean_up run () |> tap export |
55208 | 282 |
|
55205 | 283 |
val important_message = |
59172 | 284 |
if mode = Normal andalso Random.random_range 0 (atp_important_message_keep_quotient - 1) = 0 |
285 |
then extract_important_message output |
|
286 |
else "" |
|
55205 | 287 |
|
57738 | 288 |
val (used_facts, preferred_methss, message) = |
55205 | 289 |
(case outcome of |
290 |
NONE => |
|
291 |
let |
|
60924
610794dff23c
tuned signature, in accordance to sortBy in Scala;
wenzelm
parents:
59582
diff
changeset
|
292 |
val used_facts = sort_by fst (used_facts_in_atp_proof ctxt (map fst used_from) atp_proof) |
55205 | 293 |
val needs_full_types = is_typed_helper_used_in_atp_proof atp_proof |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57732
diff
changeset
|
294 |
val preferred_methss = |
57778
cf4215911f85
default on 'metis' for ATPs if preplaying is disabled
blanchet
parents:
57776
diff
changeset
|
295 |
(Metis_Method (NONE, NONE), |
72518
4be6ae020fc4
Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents:
72511
diff
changeset
|
296 |
bunches_of_proof_methods ctxt try0 smt_proofs needs_full_types |
75064 | 297 |
(if atp_proof_prefers_lifting atp_proof then liftingN else opaque_liftingN)) |
55205 | 298 |
in |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57732
diff
changeset
|
299 |
(used_facts, preferred_methss, |
55205 | 300 |
fn preplay => |
301 |
let |
|
58843 | 302 |
val _ = if verbose then writeln "Generating proof text..." else () |
57263 | 303 |
|
55205 | 304 |
fun isar_params () = |
305 |
let |
|
75064 | 306 |
val full_atp_proof = |
307 |
atp_proof_of_tstplike_proof true (perhaps (try (unprefix remote_prefix)) name) |
|
308 |
atp_problem tstplike_proof |
|
55205 | 309 |
val metis_type_enc = |
75064 | 310 |
if is_typed_helper_used_in_atp_proof full_atp_proof then |
311 |
SOME full_typesN |
|
312 |
else |
|
313 |
NONE |
|
55257 | 314 |
val metis_lam_trans = |
75064 | 315 |
if atp_proof_prefers_lifting full_atp_proof then SOME liftingN else NONE |
316 |
val full_atp_proof = |
|
317 |
full_atp_proof |
|
70931 | 318 |
|> termify_atp_proof ctxt name format type_enc pool lifted sym_tab |
72401
2783779b7dd3
removed obsolete unmaintained experimental prover Pirate
blanchet
parents:
72400
diff
changeset
|
319 |
|> local_name = spassN ? introduce_spass_skolems |
57263 | 320 |
|> factify_atp_proof (map fst used_from) hyp_ts concl_t |
55205 | 321 |
in |
57245 | 322 |
(verbose, (metis_type_enc, metis_lam_trans), preplay_timeout, compress, try0, |
75064 | 323 |
minimize, full_atp_proof, goal) |
55205 | 324 |
end |
57263 | 325 |
|
57750 | 326 |
val one_line_params = (preplay (), proof_banner mode name, subgoal, subgoal_count) |
55205 | 327 |
val num_chained = length (#facts (Proof.goal state)) |
328 |
in |
|
57738 | 329 |
proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained |
330 |
one_line_params ^ |
|
331 |
(if important_message <> "" then |
|
332 |
"\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ important_message |
|
333 |
else |
|
334 |
"") |
|
335 |
end) |
|
55205 | 336 |
end |
337 |
| SOME failure => |
|
57738 | 338 |
([], (Auto_Method (* dummy *), []), fn _ => string_of_atp_failure failure)) |
55205 | 339 |
in |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57732
diff
changeset
|
340 |
{outcome = outcome, used_facts = used_facts, used_from = used_from, |
57738 | 341 |
preferred_methss = preferred_methss, run_time = run_time, message = message} |
55205 | 342 |
end |
343 |
||
62391 | 344 |
end; |