author | blanchet |
Thu, 29 Jul 2010 22:43:46 +0200 | |
changeset 38100 | e458a0dd3dc1 |
parent 38098 | db90d313cf53 |
child 38102 | 019a49759829 |
permissions | -rw-r--r-- |
38021
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents:
38020
diff
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:
32995
diff
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:
38020
diff
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 |
|
38021
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents:
38020
diff
changeset
|
9 |
signature SLEDGEHAMMER = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
10 |
sig |
38023 | 11 |
type failure = ATP_Systems.failure |
35969 | 12 |
type relevance_override = Sledgehammer_Fact_Filter.relevance_override |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36235
diff
changeset
|
13 |
type minimize_command = Sledgehammer_Proof_Reconstruct.minimize_command |
35969 | 14 |
type params = |
15 |
{debug: bool, |
|
16 |
verbose: bool, |
|
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36064
diff
changeset
|
17 |
overlord: bool, |
35969 | 18 |
atps: string list, |
19 |
full_types: bool, |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36231
diff
changeset
|
20 |
explicit_apply: bool, |
35969 | 21 |
relevance_threshold: real, |
36922 | 22 |
relevance_convergence: real, |
36220
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents:
36184
diff
changeset
|
23 |
theory_relevant: bool option, |
36922 | 24 |
defs_relevant: bool, |
35969 | 25 |
isar_proof: bool, |
36924 | 26 |
isar_shrink_factor: int, |
35969 | 27 |
timeout: Time.time, |
28 |
minimize_timeout: Time.time} |
|
35867 | 29 |
type problem = |
35969 | 30 |
{subgoal: int, |
31 |
goal: Proof.context * (thm list * thm), |
|
32 |
relevance_override: relevance_override, |
|
38084
e2aac207d13b
"axiom_clauses" -> "axioms" (these are no longer clauses)
blanchet
parents:
38083
diff
changeset
|
33 |
axioms: (string * thm) list option} |
35867 | 34 |
type prover_result = |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
35 |
{outcome: failure option, |
35969 | 36 |
message: string, |
37926
e6ff246c0cdb
renamings + only need second component of name pool to reconstruct proofs
blanchet
parents:
37627
diff
changeset
|
37 |
pool: string Symtab.table, |
38015 | 38 |
used_thm_names: string list, |
35969 | 39 |
atp_run_time_in_msecs: int, |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
40 |
output: string, |
35969 | 41 |
proof: string, |
42 |
internal_thm_names: string Vector.vector, |
|
38083
c4b57f68ddb3
remove the "extra_clauses" business introduced in 19a5f1c8a844;
blanchet
parents:
38061
diff
changeset
|
43 |
conjecture_shape: int list list} |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38098
diff
changeset
|
44 |
type prover = params -> minimize_command -> problem -> prover_result |
35867 | 45 |
|
38023 | 46 |
val dest_dir : string Config.T |
47 |
val problem_prefix : string Config.T |
|
48 |
val measure_runtime : bool Config.T |
|
35969 | 49 |
val kill_atps: unit -> unit |
50 |
val running_atps: unit -> unit |
|
29112
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
wenzelm
parents:
28835
diff
changeset
|
51 |
val messages: int option -> unit |
38023 | 52 |
val get_prover_fun : theory -> string -> prover |
38044 | 53 |
val run_sledgehammer : |
54 |
params -> int -> relevance_override -> (string -> minimize_command) |
|
55 |
-> Proof.state -> unit |
|
38023 | 56 |
val setup : theory -> theory |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
57 |
end; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
58 |
|
38021
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents:
38020
diff
changeset
|
59 |
structure Sledgehammer : SLEDGEHAMMER = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
60 |
struct |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
61 |
|
38028 | 62 |
open ATP_Problem |
63 |
open ATP_Systems |
|
37578
9367cb36b1c4
renamed "Sledgehammer_FOL_Clauses" to "Metis_Clauses", so that Metis doesn't depend on Sledgehammer
blanchet
parents:
37577
diff
changeset
|
64 |
open Metis_Clauses |
38023 | 65 |
open Sledgehammer_Util |
36063
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36059
diff
changeset
|
66 |
open Sledgehammer_Fact_Filter |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36059
diff
changeset
|
67 |
open Sledgehammer_Proof_Reconstruct |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
68 |
|
38023 | 69 |
|
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
70 |
(** The Sledgehammer **) |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
71 |
|
37585 | 72 |
val das_Tool = "Sledgehammer" |
73 |
||
74 |
fun kill_atps () = Async_Manager.kill_threads das_Tool "ATPs" |
|
75 |
fun running_atps () = Async_Manager.running_threads das_Tool "ATPs" |
|
76 |
val messages = Async_Manager.thread_messages das_Tool "ATP" |
|
35969 | 77 |
|
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36235
diff
changeset
|
78 |
(** problems, results, provers, etc. **) |
35969 | 79 |
|
80 |
type params = |
|
81 |
{debug: bool, |
|
82 |
verbose: bool, |
|
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36064
diff
changeset
|
83 |
overlord: bool, |
35969 | 84 |
atps: string list, |
85 |
full_types: bool, |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36231
diff
changeset
|
86 |
explicit_apply: bool, |
35969 | 87 |
relevance_threshold: real, |
36922 | 88 |
relevance_convergence: real, |
36220
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents:
36184
diff
changeset
|
89 |
theory_relevant: bool option, |
36922 | 90 |
defs_relevant: bool, |
35969 | 91 |
isar_proof: bool, |
36924 | 92 |
isar_shrink_factor: int, |
35969 | 93 |
timeout: Time.time, |
94 |
minimize_timeout: Time.time} |
|
35867 | 95 |
|
96 |
type problem = |
|
35969 | 97 |
{subgoal: int, |
98 |
goal: Proof.context * (thm list * thm), |
|
99 |
relevance_override: relevance_override, |
|
38084
e2aac207d13b
"axiom_clauses" -> "axioms" (these are no longer clauses)
blanchet
parents:
38083
diff
changeset
|
100 |
axioms: (string * thm) list option} |
35867 | 101 |
|
102 |
type prover_result = |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
103 |
{outcome: failure option, |
35969 | 104 |
message: string, |
37926
e6ff246c0cdb
renamings + only need second component of name pool to reconstruct proofs
blanchet
parents:
37627
diff
changeset
|
105 |
pool: string Symtab.table, |
38015 | 106 |
used_thm_names: string list, |
35969 | 107 |
atp_run_time_in_msecs: int, |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
108 |
output: string, |
35969 | 109 |
proof: string, |
110 |
internal_thm_names: string Vector.vector, |
|
38083
c4b57f68ddb3
remove the "extra_clauses" business introduced in 19a5f1c8a844;
blanchet
parents:
38061
diff
changeset
|
111 |
conjecture_shape: int list list} |
35867 | 112 |
|
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38098
diff
changeset
|
113 |
type prover = params -> minimize_command -> problem -> prover_result |
35867 | 114 |
|
38023 | 115 |
(* configuration attributes *) |
116 |
||
117 |
val (dest_dir, dest_dir_setup) = Attrib.config_string "atp_dest_dir" (K ""); |
|
118 |
(*Empty string means create files in Isabelle's temporary files directory.*) |
|
119 |
||
120 |
val (problem_prefix, problem_prefix_setup) = |
|
121 |
Attrib.config_string "atp_problem_prefix" (K "prob"); |
|
122 |
||
123 |
val (measure_runtime, measure_runtime_setup) = |
|
124 |
Attrib.config_bool "atp_measure_runtime" (K false); |
|
28484 | 125 |
|
38023 | 126 |
fun with_path cleanup after f path = |
127 |
Exn.capture f path |
|
128 |
|> tap (fn _ => cleanup path) |
|
129 |
|> Exn.release |
|
130 |
|> tap (after path) |
|
131 |
||
132 |
(* Splits by the first possible of a list of delimiters. *) |
|
133 |
fun extract_proof delims output = |
|
134 |
case pairself (find_first (fn s => String.isSubstring s output)) |
|
135 |
(ListPair.unzip delims) of |
|
136 |
(SOME begin_delim, SOME end_delim) => |
|
137 |
(output |> first_field begin_delim |> the |> snd |
|
138 |
|> first_field end_delim |> the |> fst |
|
139 |
|> first_field "\n" |> the |> snd |
|
140 |
handle Option.Option => "") |
|
141 |
| _ => "" |
|
28484 | 142 |
|
38023 | 143 |
fun extract_proof_and_outcome complete res_code proof_delims known_failures |
144 |
output = |
|
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38044
diff
changeset
|
145 |
case known_failure_in_output output known_failures of |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38044
diff
changeset
|
146 |
NONE => (case extract_proof proof_delims output of |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38044
diff
changeset
|
147 |
"" => ("", SOME MalformedOutput) |
38023 | 148 |
| proof => if res_code = 0 then (proof, NONE) |
149 |
else ("", SOME UnknownError)) |
|
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38044
diff
changeset
|
150 |
| SOME failure => |
38023 | 151 |
("", SOME (if failure = IncompleteUnprovable andalso complete then |
152 |
Unprovable |
|
153 |
else |
|
154 |
failure)) |
|
28582 | 155 |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
156 |
|
38023 | 157 |
(* Clause preparation *) |
158 |
||
159 |
datatype fol_formula = |
|
160 |
FOLFormula of {formula_name: string, |
|
161 |
kind: kind, |
|
162 |
combformula: (name, combterm) formula, |
|
163 |
ctypes_sorts: typ list} |
|
164 |
||
165 |
fun mk_anot phi = AConn (ANot, [phi]) |
|
166 |
fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2]) |
|
167 |
fun mk_ahorn [] phi = phi |
|
168 |
| mk_ahorn (phi :: phis) psi = |
|
169 |
AConn (AImplies, [fold (mk_aconn AAnd) phis phi, psi]) |
|
170 |
||
171 |
(* ### FIXME: reintroduce |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
172 |
fun make_formula_table xs = |
38023 | 173 |
fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty |
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
174 |
(* Remove existing axioms from the conjecture, as this can |
38023 | 175 |
dramatically boost an ATP's performance (for some reason). *) |
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
176 |
fun subtract_cls axioms = |
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
177 |
filter_out (Termtab.defined (make_formula_table axioms) o prop_of) |
38023 | 178 |
*) |
179 |
||
180 |
fun combformula_for_prop thy = |
|
181 |
let |
|
182 |
val do_term = combterm_from_term thy |
|
183 |
fun do_quant bs q s T t' = |
|
184 |
do_formula ((s, T) :: bs) t' |
|
185 |
#>> (fn phi => AQuant (q, [`make_bound_var s], phi)) |
|
186 |
and do_conn bs c t1 t2 = |
|
187 |
do_formula bs t1 ##>> do_formula bs t2 |
|
188 |
#>> (fn (phi1, phi2) => AConn (c, [phi1, phi2])) |
|
189 |
and do_formula bs t = |
|
190 |
case t of |
|
191 |
@{const Not} $ t1 => |
|
192 |
do_formula bs t1 #>> (fn phi => AConn (ANot, [phi])) |
|
193 |
| Const (@{const_name All}, _) $ Abs (s, T, t') => |
|
194 |
do_quant bs AForall s T t' |
|
195 |
| Const (@{const_name Ex}, _) $ Abs (s, T, t') => |
|
196 |
do_quant bs AExists s T t' |
|
197 |
| @{const "op &"} $ t1 $ t2 => do_conn bs AAnd t1 t2 |
|
198 |
| @{const "op |"} $ t1 $ t2 => do_conn bs AOr t1 t2 |
|
199 |
| @{const "op -->"} $ t1 $ t2 => do_conn bs AImplies t1 t2 |
|
200 |
| Const (@{const_name "op ="}, Type (_, [@{typ bool}, _])) $ t1 $ t2 => |
|
201 |
do_conn bs AIff t1 t2 |
|
202 |
| _ => (fn ts => do_term bs (Envir.eta_contract t) |
|
38034 | 203 |
|>> AAtom ||> union (op =) ts) |
38023 | 204 |
in do_formula [] end |
205 |
||
206 |
(* Converts an elim-rule into an equivalent theorem that does not have the |
|
207 |
predicate variable. Leaves other theorems unchanged. We simply instantiate |
|
208 |
the conclusion variable to False. (Cf. "transform_elim_term" in |
|
209 |
"ATP_Systems".) *) |
|
210 |
(* FIXME: test! *) |
|
211 |
fun transform_elim_term t = |
|
212 |
case Logic.strip_imp_concl t of |
|
213 |
@{const Trueprop} $ Var (z, @{typ bool}) => |
|
214 |
subst_Vars [(z, @{const True})] t |
|
215 |
| Var (z, @{typ prop}) => subst_Vars [(z, @{prop True})] t |
|
216 |
| _ => t |
|
217 |
||
218 |
(* Removes the lambdas from an equation of the form "t = (%x. u)". |
|
219 |
(Cf. "extensionalize_theorem" in "Clausifier".) *) |
|
220 |
fun extensionalize_term t = |
|
221 |
let |
|
222 |
fun aux j (Const (@{const_name "op ="}, Type (_, [Type (_, [_, T']), _])) |
|
223 |
$ t2 $ Abs (s, var_T, t')) = |
|
224 |
let val var_t = Var (("x", j), var_T) in |
|
225 |
Const (@{const_name "op ="}, T' --> T' --> HOLogic.boolT) |
|
226 |
$ betapply (t2, var_t) $ subst_bound (var_t, t') |
|
227 |
|> aux (j + 1) |
|
228 |
end |
|
229 |
| aux _ t = t |
|
230 |
in aux (maxidx_of_term t + 1) t end |
|
231 |
||
38089
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
232 |
fun presimplify_term thy = |
38091 | 233 |
Skip_Proof.make_thm thy |
38089
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
234 |
#> Meson.presimplify |
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
235 |
#> prop_of |
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
236 |
|
38098
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
237 |
(* Freshness almost guaranteed! *) |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
238 |
fun concealed_bound_name j = das_Tool ^ Int.toString j |
38023 | 239 |
fun conceal_bounds Ts t = |
240 |
subst_bounds (map (Free o apfst concealed_bound_name) |
|
241 |
(length Ts - 1 downto 0 ~~ rev Ts), t) |
|
242 |
fun reveal_bounds Ts = |
|
243 |
subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j)) |
|
244 |
(0 upto length Ts - 1 ~~ Ts)) |
|
245 |
||
246 |
fun introduce_combinators_in_term ctxt kind t = |
|
247 |
let |
|
248 |
val thy = ProofContext.theory_of ctxt |
|
249 |
fun aux Ts t = |
|
250 |
case t of |
|
251 |
@{const Not} $ t1 => @{const Not} $ aux Ts t1 |
|
252 |
| (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') => |
|
253 |
t0 $ Abs (s, T, aux (T :: Ts) t') |
|
254 |
| (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') => |
|
255 |
t0 $ Abs (s, T, aux (T :: Ts) t') |
|
256 |
| (t0 as @{const "op &"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 |
|
257 |
| (t0 as @{const "op |"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 |
|
258 |
| (t0 as @{const "op -->"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 |
|
259 |
| (t0 as Const (@{const_name "op ="}, Type (_, [@{typ bool}, _]))) |
|
260 |
$ t1 $ t2 => |
|
261 |
t0 $ aux Ts t1 $ aux Ts t2 |
|
262 |
| _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then |
|
263 |
t |
|
264 |
else |
|
265 |
let |
|
266 |
val t = t |> conceal_bounds Ts |
|
267 |
|> Envir.eta_contract |
|
268 |
val ([t], ctxt') = Variable.import_terms true [t] ctxt |
|
269 |
in |
|
270 |
t |> cterm_of thy |
|
271 |
|> Clausifier.introduce_combinators_in_cterm |
|
272 |
|> singleton (Variable.export ctxt' ctxt) |
|
273 |
|> prop_of |> Logic.dest_equals |> snd |
|
274 |
|> reveal_bounds Ts |
|
275 |
end |
|
276 |
in t |> not (Meson.is_fol_term thy t) ? aux [] end |
|
277 |
handle THM _ => |
|
278 |
(* A type variable of sort "{}" will make abstraction fail. *) |
|
279 |
case kind of |
|
280 |
Axiom => HOLogic.true_const |
|
281 |
| Conjecture => HOLogic.false_const |
|
282 |
||
38098
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
283 |
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
284 |
same in Sledgehammer to prevent the discovery of unreplable proofs. *) |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
285 |
fun freeze_term t = |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
286 |
let |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
287 |
fun aux (t $ u) = aux t $ aux u |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
288 |
| aux (Abs (s, T, t)) = Abs (s, T, aux t) |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
289 |
| aux (Var ((s, i), T)) = |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
290 |
Free (das_Tool ^ "_" ^ s ^ "_" ^ string_of_int i, T) |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
291 |
| aux t = t |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
292 |
in t |> exists_subterm is_Var t ? aux end |
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
293 |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
294 |
(* making axiom and conjecture formulas *) |
38089
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
295 |
fun make_formula ctxt presimp (formula_name, kind, t) = |
38023 | 296 |
let |
297 |
val thy = ProofContext.theory_of ctxt |
|
298 |
val t = t |> transform_elim_term |
|
299 |
|> Object_Logic.atomize_term thy |
|
38091 | 300 |
val t = t |> fastype_of t = HOLogic.boolT ? HOLogic.mk_Trueprop |
38023 | 301 |
|> extensionalize_term |
38089
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
302 |
|> presimp ? presimplify_term thy |
38091 | 303 |
|> perhaps (try (HOLogic.dest_Trueprop)) |
38023 | 304 |
|> introduce_combinators_in_term ctxt kind |
38098
db90d313cf53
handle schematic vars the same way in Sledgehammer as in Metis, to avoid unreplayable proofs
blanchet
parents:
38092
diff
changeset
|
305 |
|> kind = Conjecture ? freeze_term |
38023 | 306 |
val (combformula, ctypes_sorts) = combformula_for_prop thy t [] |
307 |
in |
|
308 |
FOLFormula {formula_name = formula_name, combformula = combformula, |
|
309 |
kind = kind, ctypes_sorts = ctypes_sorts} |
|
310 |
end |
|
311 |
||
38089
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
312 |
fun make_axiom ctxt presimp (name, th) = |
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
313 |
(name, make_formula ctxt presimp (name, Axiom, prop_of th)) |
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
314 |
fun make_conjectures ctxt ts = |
38089
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
315 |
map2 (fn j => fn t => make_formula ctxt true (Int.toString j, Conjecture, t)) |
38023 | 316 |
(0 upto length ts - 1) ts |
317 |
||
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
318 |
(** Helper facts **) |
38023 | 319 |
|
320 |
fun count_combterm (CombConst ((s, _), _, _)) = |
|
321 |
Symtab.map_entry s (Integer.add 1) |
|
322 |
| count_combterm (CombVar _) = I |
|
323 |
| count_combterm (CombApp (t1, t2)) = fold count_combterm [t1, t2] |
|
324 |
fun count_combformula (AQuant (_, _, phi)) = count_combformula phi |
|
325 |
| count_combformula (AConn (_, phis)) = fold count_combformula phis |
|
38034 | 326 |
| count_combformula (AAtom tm) = count_combterm tm |
38023 | 327 |
fun count_fol_formula (FOLFormula {combformula, ...}) = |
328 |
count_combformula combformula |
|
329 |
||
330 |
val optional_helpers = |
|
331 |
[(["c_COMBI", "c_COMBK"], @{thms COMBI_def COMBK_def}), |
|
332 |
(["c_COMBB", "c_COMBC"], @{thms COMBB_def COMBC_def}), |
|
333 |
(["c_COMBS"], @{thms COMBS_def})] |
|
334 |
val optional_typed_helpers = |
|
335 |
[(["c_True", "c_False"], @{thms True_or_False}), |
|
336 |
(["c_If"], @{thms if_True if_False True_or_False})] |
|
337 |
val mandatory_helpers = @{thms fequal_imp_equal equal_imp_fequal} |
|
338 |
||
339 |
val init_counters = |
|
340 |
Symtab.make (maps (maps (map (rpair 0) o fst)) |
|
341 |
[optional_helpers, optional_typed_helpers]) |
|
342 |
||
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
343 |
fun get_helper_facts ctxt is_FO full_types conjectures axioms = |
38023 | 344 |
let |
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
345 |
val ct = fold (fold count_fol_formula) [conjectures, axioms] init_counters |
38023 | 346 |
fun is_needed c = the (Symtab.lookup ct c) > 0 |
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
347 |
in |
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
348 |
(optional_helpers |
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
349 |
|> full_types ? append optional_typed_helpers |
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
350 |
|> maps (fn (ss, ths) => |
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
351 |
if exists is_needed ss then map (`Thm.get_name_hint) ths |
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
352 |
else [])) @ |
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
353 |
(if is_FO then [] else map (`Thm.get_name_hint) mandatory_helpers) |
38089
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
354 |
|> map (snd o make_axiom ctxt false) |
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
355 |
end |
38023 | 356 |
|
38086 | 357 |
fun meta_not t = @{const "==>"} $ t $ @{prop False} |
38023 | 358 |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
359 |
fun prepare_formulas ctxt full_types hyp_ts concl_t axcls = |
38023 | 360 |
let |
361 |
val thy = ProofContext.theory_of ctxt |
|
362 |
val goal_t = Logic.list_implies (hyp_ts, concl_t) |
|
363 |
val is_FO = Meson.is_fol_term thy goal_t |
|
38083
c4b57f68ddb3
remove the "extra_clauses" business introduced in 19a5f1c8a844;
blanchet
parents:
38061
diff
changeset
|
364 |
val axtms = map (prop_of o snd) axcls |
38023 | 365 |
val subs = tfree_classes_of_terms [goal_t] |
366 |
val supers = tvar_classes_of_terms axtms |
|
367 |
val tycons = type_consts_of_terms thy (goal_t :: axtms) |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
368 |
(* TFrees in the conjecture; TVars in the axioms *) |
38086 | 369 |
val conjectures = map meta_not hyp_ts @ [concl_t] |> make_conjectures ctxt |
38089
ed65a0777e10
perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents:
38088
diff
changeset
|
370 |
val (clnames, axioms) = ListPair.unzip (map (make_axiom ctxt true) axcls) |
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
371 |
val helper_facts = get_helper_facts ctxt is_FO full_types conjectures axioms |
38023 | 372 |
val (supers', arity_clauses) = make_arity_clauses thy tycons supers |
373 |
val class_rel_clauses = make_class_rel_clauses thy subs supers' |
|
374 |
in |
|
375 |
(Vector.fromList clnames, |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
376 |
(conjectures, axioms, helper_facts, class_rel_clauses, arity_clauses)) |
38023 | 377 |
end |
378 |
||
379 |
fun wrap_type ty t = ATerm ((type_wrapper_name, type_wrapper_name), [ty, t]) |
|
380 |
||
381 |
fun fo_term_for_combtyp (CombTVar name) = ATerm (name, []) |
|
382 |
| fo_term_for_combtyp (CombTFree name) = ATerm (name, []) |
|
383 |
| fo_term_for_combtyp (CombType (name, tys)) = |
|
384 |
ATerm (name, map fo_term_for_combtyp tys) |
|
385 |
||
386 |
fun fo_literal_for_type_literal (TyLitVar (class, name)) = |
|
387 |
(true, ATerm (class, [ATerm (name, [])])) |
|
388 |
| fo_literal_for_type_literal (TyLitFree (class, name)) = |
|
389 |
(true, ATerm (class, [ATerm (name, [])])) |
|
390 |
||
38034 | 391 |
fun formula_for_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot |
38023 | 392 |
|
393 |
fun fo_term_for_combterm full_types = |
|
394 |
let |
|
395 |
fun aux top_level u = |
|
396 |
let |
|
397 |
val (head, args) = strip_combterm_comb u |
|
398 |
val (x, ty_args) = |
|
399 |
case head of |
|
38087
dddb8ba3a1ce
generate correct names for "$true" and "$false";
blanchet
parents:
38086
diff
changeset
|
400 |
CombConst (name as (s, s'), _, ty_args) => |
38088
a9847fb539dd
fix bug with "=" vs. "fequal" introduced by last change (dddb8ba3a1ce)
blanchet
parents:
38087
diff
changeset
|
401 |
if s = "equal" then |
a9847fb539dd
fix bug with "=" vs. "fequal" introduced by last change (dddb8ba3a1ce)
blanchet
parents:
38087
diff
changeset
|
402 |
(if top_level andalso length args = 2 then name |
a9847fb539dd
fix bug with "=" vs. "fequal" introduced by last change (dddb8ba3a1ce)
blanchet
parents:
38087
diff
changeset
|
403 |
else ("c_fequal", @{const_name fequal}), []) |
a9847fb539dd
fix bug with "=" vs. "fequal" introduced by last change (dddb8ba3a1ce)
blanchet
parents:
38087
diff
changeset
|
404 |
else if top_level then |
a9847fb539dd
fix bug with "=" vs. "fequal" introduced by last change (dddb8ba3a1ce)
blanchet
parents:
38087
diff
changeset
|
405 |
case s of |
a9847fb539dd
fix bug with "=" vs. "fequal" introduced by last change (dddb8ba3a1ce)
blanchet
parents:
38087
diff
changeset
|
406 |
"c_False" => (("$false", s'), []) |
a9847fb539dd
fix bug with "=" vs. "fequal" introduced by last change (dddb8ba3a1ce)
blanchet
parents:
38087
diff
changeset
|
407 |
| "c_True" => (("$true", s'), []) |
a9847fb539dd
fix bug with "=" vs. "fequal" introduced by last change (dddb8ba3a1ce)
blanchet
parents:
38087
diff
changeset
|
408 |
| _ => (name, if full_types then [] else ty_args) |
38023 | 409 |
else |
410 |
(name, if full_types then [] else ty_args) |
|
411 |
| CombVar (name, _) => (name, []) |
|
412 |
| CombApp _ => raise Fail "impossible \"CombApp\"" |
|
413 |
val t = ATerm (x, map fo_term_for_combtyp ty_args @ |
|
414 |
map (aux false) args) |
|
415 |
in |
|
416 |
if full_types then wrap_type (fo_term_for_combtyp (combtyp_of u)) t else t |
|
417 |
end |
|
418 |
in aux true end |
|
419 |
||
420 |
fun formula_for_combformula full_types = |
|
421 |
let |
|
422 |
fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi) |
|
423 |
| aux (AConn (c, phis)) = AConn (c, map aux phis) |
|
38034 | 424 |
| aux (AAtom tm) = AAtom (fo_term_for_combterm full_types tm) |
38023 | 425 |
in aux end |
426 |
||
427 |
fun formula_for_axiom full_types (FOLFormula {combformula, ctypes_sorts, ...}) = |
|
428 |
mk_ahorn (map (formula_for_fo_literal o fo_literal_for_type_literal) |
|
429 |
(type_literals_for_types ctypes_sorts)) |
|
430 |
(formula_for_combformula full_types combformula) |
|
431 |
||
432 |
fun problem_line_for_axiom full_types |
|
433 |
(formula as FOLFormula {formula_name, kind, ...}) = |
|
434 |
Fof (axiom_prefix ^ ascii_of formula_name, kind, |
|
435 |
formula_for_axiom full_types formula) |
|
436 |
||
437 |
fun problem_line_for_class_rel_clause |
|
438 |
(ClassRelClause {axiom_name, subclass, superclass, ...}) = |
|
439 |
let val ty_arg = ATerm (("T", "T"), []) in |
|
440 |
Fof (ascii_of axiom_name, Axiom, |
|
38034 | 441 |
AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])), |
442 |
AAtom (ATerm (superclass, [ty_arg]))])) |
|
38023 | 443 |
end |
444 |
||
445 |
fun fo_literal_for_arity_literal (TConsLit (c, t, args)) = |
|
446 |
(true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)])) |
|
447 |
| fo_literal_for_arity_literal (TVarLit (c, sort)) = |
|
448 |
(false, ATerm (c, [ATerm (sort, [])])) |
|
449 |
||
450 |
fun problem_line_for_arity_clause |
|
451 |
(ArityClause {axiom_name, conclLit, premLits, ...}) = |
|
452 |
Fof (arity_clause_prefix ^ ascii_of axiom_name, Axiom, |
|
453 |
mk_ahorn (map (formula_for_fo_literal o apfst not |
|
454 |
o fo_literal_for_arity_literal) premLits) |
|
455 |
(formula_for_fo_literal |
|
456 |
(fo_literal_for_arity_literal conclLit))) |
|
457 |
||
458 |
fun problem_line_for_conjecture full_types |
|
459 |
(FOLFormula {formula_name, kind, combformula, ...}) = |
|
460 |
Fof (conjecture_prefix ^ formula_name, kind, |
|
461 |
formula_for_combformula full_types combformula) |
|
462 |
||
463 |
fun free_type_literals_for_conjecture (FOLFormula {ctypes_sorts, ...}) = |
|
464 |
map fo_literal_for_type_literal (type_literals_for_types ctypes_sorts) |
|
465 |
||
466 |
fun problem_line_for_free_type lit = |
|
467 |
Fof (tfrees_name, Conjecture, mk_anot (formula_for_fo_literal lit)) |
|
468 |
fun problem_lines_for_free_types conjectures = |
|
469 |
let |
|
470 |
val litss = map free_type_literals_for_conjecture conjectures |
|
471 |
val lits = fold (union (op =)) litss [] |
|
472 |
in map problem_line_for_free_type lits end |
|
473 |
||
474 |
(** "hBOOL" and "hAPP" **) |
|
475 |
||
476 |
type const_info = {min_arity: int, max_arity: int, sub_level: bool} |
|
477 |
||
478 |
fun consider_term top_level (ATerm ((s, _), ts)) = |
|
479 |
(if is_tptp_variable s then |
|
480 |
I |
|
481 |
else |
|
482 |
let val n = length ts in |
|
483 |
Symtab.map_default |
|
484 |
(s, {min_arity = n, max_arity = 0, sub_level = false}) |
|
485 |
(fn {min_arity, max_arity, sub_level} => |
|
486 |
{min_arity = Int.min (n, min_arity), |
|
487 |
max_arity = Int.max (n, max_arity), |
|
488 |
sub_level = sub_level orelse not top_level}) |
|
489 |
end) |
|
490 |
#> fold (consider_term (top_level andalso s = type_wrapper_name)) ts |
|
491 |
fun consider_formula (AQuant (_, _, phi)) = consider_formula phi |
|
492 |
| consider_formula (AConn (_, phis)) = fold consider_formula phis |
|
38034 | 493 |
| consider_formula (AAtom tm) = consider_term true tm |
38023 | 494 |
|
495 |
fun consider_problem_line (Fof (_, _, phi)) = consider_formula phi |
|
496 |
fun consider_problem problem = fold (fold consider_problem_line o snd) problem |
|
497 |
||
498 |
fun const_table_for_problem explicit_apply problem = |
|
499 |
if explicit_apply then NONE |
|
500 |
else SOME (Symtab.empty |> consider_problem problem) |
|
501 |
||
502 |
val tc_fun = make_fixed_type_const @{type_name fun} |
|
503 |
||
504 |
fun min_arity_of thy full_types NONE s = |
|
505 |
(if s = "equal" orelse s = type_wrapper_name orelse |
|
506 |
String.isPrefix type_const_prefix s orelse |
|
507 |
String.isPrefix class_prefix s then |
|
508 |
16383 (* large number *) |
|
509 |
else if full_types then |
|
510 |
0 |
|
511 |
else case strip_prefix_and_undo_ascii const_prefix s of |
|
512 |
SOME s' => num_type_args thy (invert_const s') |
|
513 |
| NONE => 0) |
|
514 |
| min_arity_of _ _ (SOME the_const_tab) s = |
|
515 |
case Symtab.lookup the_const_tab s of |
|
516 |
SOME ({min_arity, ...} : const_info) => min_arity |
|
517 |
| NONE => 0 |
|
518 |
||
519 |
fun full_type_of (ATerm ((s, _), [ty, _])) = |
|
520 |
if s = type_wrapper_name then ty else raise Fail "expected type wrapper" |
|
521 |
| full_type_of _ = raise Fail "expected type wrapper" |
|
522 |
||
523 |
fun list_hAPP_rev _ t1 [] = t1 |
|
524 |
| list_hAPP_rev NONE t1 (t2 :: ts2) = |
|
525 |
ATerm (`I "hAPP", [list_hAPP_rev NONE t1 ts2, t2]) |
|
526 |
| list_hAPP_rev (SOME ty) t1 (t2 :: ts2) = |
|
527 |
let val ty' = ATerm (`make_fixed_type_const @{type_name fun}, |
|
528 |
[full_type_of t2, ty]) in |
|
529 |
ATerm (`I "hAPP", [wrap_type ty' (list_hAPP_rev (SOME ty') t1 ts2), t2]) |
|
530 |
end |
|
531 |
||
532 |
fun repair_applications_in_term thy full_types const_tab = |
|
533 |
let |
|
534 |
fun aux opt_ty (ATerm (name as (s, _), ts)) = |
|
535 |
if s = type_wrapper_name then |
|
536 |
case ts of |
|
537 |
[t1, t2] => ATerm (name, [aux NONE t1, aux (SOME t1) t2]) |
|
538 |
| _ => raise Fail "malformed type wrapper" |
|
539 |
else |
|
540 |
let |
|
541 |
val ts = map (aux NONE) ts |
|
542 |
val (ts1, ts2) = chop (min_arity_of thy full_types const_tab s) ts |
|
543 |
in list_hAPP_rev opt_ty (ATerm (name, ts1)) (rev ts2) end |
|
544 |
in aux NONE end |
|
545 |
||
546 |
fun boolify t = ATerm (`I "hBOOL", [t]) |
|
547 |
||
548 |
(* True if the constant ever appears outside of the top-level position in |
|
549 |
literals, or if it appears with different arities (e.g., because of different |
|
550 |
type instantiations). If false, the constant always receives all of its |
|
551 |
arguments and is used as a predicate. *) |
|
552 |
fun is_predicate NONE s = |
|
553 |
s = "equal" orelse String.isPrefix type_const_prefix s orelse |
|
554 |
String.isPrefix class_prefix s |
|
555 |
| is_predicate (SOME the_const_tab) s = |
|
556 |
case Symtab.lookup the_const_tab s of |
|
557 |
SOME {min_arity, max_arity, sub_level} => |
|
558 |
not sub_level andalso min_arity = max_arity |
|
559 |
| NONE => false |
|
560 |
||
561 |
fun repair_predicates_in_term const_tab (t as ATerm ((s, _), ts)) = |
|
562 |
if s = type_wrapper_name then |
|
563 |
case ts of |
|
564 |
[_, t' as ATerm ((s', _), _)] => |
|
565 |
if is_predicate const_tab s' then t' else boolify t |
|
566 |
| _ => raise Fail "malformed type wrapper" |
|
567 |
else |
|
568 |
t |> not (is_predicate const_tab s) ? boolify |
|
569 |
||
570 |
fun close_universally phi = |
|
571 |
let |
|
572 |
fun term_vars bounds (ATerm (name as (s, _), tms)) = |
|
573 |
(is_tptp_variable s andalso not (member (op =) bounds name)) |
|
574 |
? insert (op =) name |
|
575 |
#> fold (term_vars bounds) tms |
|
576 |
fun formula_vars bounds (AQuant (q, xs, phi)) = |
|
577 |
formula_vars (xs @ bounds) phi |
|
578 |
| formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis |
|
38034 | 579 |
| formula_vars bounds (AAtom tm) = term_vars bounds tm |
38023 | 580 |
in |
581 |
case formula_vars [] phi [] of [] => phi | xs => AQuant (AForall, xs, phi) |
|
582 |
end |
|
583 |
||
584 |
fun repair_formula thy explicit_forall full_types const_tab = |
|
585 |
let |
|
586 |
fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi) |
|
587 |
| aux (AConn (c, phis)) = AConn (c, map aux phis) |
|
38034 | 588 |
| aux (AAtom tm) = |
589 |
AAtom (tm |> repair_applications_in_term thy full_types const_tab |
|
38023 | 590 |
|> repair_predicates_in_term const_tab) |
591 |
in aux #> explicit_forall ? close_universally end |
|
592 |
||
593 |
fun repair_problem_line thy explicit_forall full_types const_tab |
|
594 |
(Fof (ident, kind, phi)) = |
|
595 |
Fof (ident, kind, repair_formula thy explicit_forall full_types const_tab phi) |
|
596 |
fun repair_problem_with_const_table thy = |
|
597 |
map o apsnd o map ooo repair_problem_line thy |
|
598 |
||
599 |
fun repair_problem thy explicit_forall full_types explicit_apply problem = |
|
600 |
repair_problem_with_const_table thy explicit_forall full_types |
|
601 |
(const_table_for_problem explicit_apply problem) problem |
|
602 |
||
603 |
fun write_tptp_file thy readable_names explicit_forall full_types explicit_apply |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
604 |
file (conjectures, axioms, helper_facts, class_rel_clauses, |
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
605 |
arity_clauses) = |
38023 | 606 |
let |
38084
e2aac207d13b
"axiom_clauses" -> "axioms" (these are no longer clauses)
blanchet
parents:
38083
diff
changeset
|
607 |
val axiom_lines = map (problem_line_for_axiom full_types) axioms |
38023 | 608 |
val class_rel_lines = |
609 |
map problem_line_for_class_rel_clause class_rel_clauses |
|
610 |
val arity_lines = map problem_line_for_arity_clause arity_clauses |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
611 |
val helper_lines = map (problem_line_for_axiom full_types) helper_facts |
38023 | 612 |
val conjecture_lines = |
613 |
map (problem_line_for_conjecture full_types) conjectures |
|
614 |
val tfree_lines = problem_lines_for_free_types conjectures |
|
615 |
(* Reordering these might or might not confuse the proof reconstruction |
|
616 |
code or the SPASS Flotter hack. *) |
|
617 |
val problem = |
|
618 |
[("Relevant facts", axiom_lines), |
|
619 |
("Class relationships", class_rel_lines), |
|
620 |
("Arity declarations", arity_lines), |
|
621 |
("Helper facts", helper_lines), |
|
622 |
("Conjectures", conjecture_lines), |
|
623 |
("Type variables", tfree_lines)] |
|
624 |
|> repair_problem thy explicit_forall full_types explicit_apply |
|
625 |
val (problem, pool) = nice_tptp_problem readable_names problem |
|
626 |
val conjecture_offset = |
|
627 |
length axiom_lines + length class_rel_lines + length arity_lines |
|
628 |
+ length helper_lines |
|
629 |
val _ = File.write_list file (strings_for_tptp_problem problem) |
|
630 |
in |
|
631 |
(case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty, |
|
632 |
conjecture_offset) |
|
633 |
end |
|
634 |
||
635 |
fun extract_clause_sequence output = |
|
636 |
let |
|
637 |
val tokens_of = String.tokens (not o Char.isAlphaNum) |
|
638 |
fun extract_num ("clause" :: (ss as _ :: _)) = |
|
639 |
Int.fromString (List.last ss) |
|
640 |
| extract_num _ = NONE |
|
641 |
in output |> split_lines |> map_filter (extract_num o tokens_of) end |
|
642 |
||
643 |
val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation" |
|
644 |
||
645 |
val parse_clause_formula_pair = |
|
646 |
$$ "(" |-- scan_integer --| $$ "," -- Symbol.scan_id --| $$ ")" |
|
647 |
--| Scan.option ($$ ",") |
|
648 |
val parse_clause_formula_relation = |
|
649 |
Scan.this_string set_ClauseFormulaRelationN |-- $$ "(" |
|
650 |
|-- Scan.repeat parse_clause_formula_pair |
|
651 |
val extract_clause_formula_relation = |
|
652 |
Substring.full |
|
653 |
#> Substring.position set_ClauseFormulaRelationN |
|
654 |
#> snd #> Substring.string #> strip_spaces #> explode |
|
655 |
#> parse_clause_formula_relation #> fst |
|
656 |
||
657 |
fun repair_conjecture_shape_and_theorem_names output conjecture_shape |
|
658 |
thm_names = |
|
659 |
if String.isSubstring set_ClauseFormulaRelationN output then |
|
660 |
(* This is a hack required for keeping track of axioms after they have been |
|
661 |
clausified by SPASS's Flotter tool. The "SPASS_TPTP" script is also part |
|
662 |
of this hack. *) |
|
663 |
let |
|
38040
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents:
38039
diff
changeset
|
664 |
val j0 = hd (hd conjecture_shape) |
38023 | 665 |
val seq = extract_clause_sequence output |
666 |
val name_map = extract_clause_formula_relation output |
|
667 |
fun renumber_conjecture j = |
|
668 |
AList.find (op =) name_map (conjecture_prefix ^ Int.toString (j - j0)) |
|
38040
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents:
38039
diff
changeset
|
669 |
|> map (fn s => find_index (curry (op =) s) seq + 1) |
38023 | 670 |
in |
38040
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents:
38039
diff
changeset
|
671 |
(conjecture_shape |> map (maps renumber_conjecture), |
38023 | 672 |
seq |> map (the o AList.lookup (op =) name_map) |
673 |
|> map (fn s => case try (unprefix axiom_prefix) s of |
|
674 |
SOME s' => undo_ascii_of s' |
|
675 |
| NONE => "") |
|
676 |
|> Vector.fromList) |
|
677 |
end |
|
678 |
else |
|
679 |
(conjecture_shape, thm_names) |
|
680 |
||
681 |
||
682 |
(* generic TPTP-based provers *) |
|
683 |
||
684 |
fun prover_fun name |
|
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38091
diff
changeset
|
685 |
{exec, required_execs, arguments, proof_delims, known_failures, |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38091
diff
changeset
|
686 |
max_new_relevant_facts_per_iter, prefers_theory_relevant, |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38091
diff
changeset
|
687 |
explicit_forall} |
38023 | 688 |
({debug, overlord, full_types, explicit_apply, relevance_threshold, |
689 |
relevance_convergence, theory_relevant, defs_relevant, isar_proof, |
|
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38098
diff
changeset
|
690 |
isar_shrink_factor, timeout, ...} : params) |
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38098
diff
changeset
|
691 |
minimize_command |
38084
e2aac207d13b
"axiom_clauses" -> "axioms" (these are no longer clauses)
blanchet
parents:
38083
diff
changeset
|
692 |
({subgoal, goal, relevance_override, axioms} : problem) = |
38023 | 693 |
let |
694 |
val (ctxt, (_, th)) = goal; |
|
695 |
val thy = ProofContext.theory_of ctxt |
|
696 |
(* ### FIXME: (1) preprocessing for "if" etc. *) |
|
697 |
val (params, hyp_ts, concl_t) = strip_subgoal th subgoal |
|
38084
e2aac207d13b
"axiom_clauses" -> "axioms" (these are no longer clauses)
blanchet
parents:
38083
diff
changeset
|
698 |
val the_axioms = |
e2aac207d13b
"axiom_clauses" -> "axioms" (these are no longer clauses)
blanchet
parents:
38083
diff
changeset
|
699 |
case axioms of |
38083
c4b57f68ddb3
remove the "extra_clauses" business introduced in 19a5f1c8a844;
blanchet
parents:
38061
diff
changeset
|
700 |
SOME axioms => axioms |
38023 | 701 |
| NONE => relevant_facts full_types relevance_threshold |
702 |
relevance_convergence defs_relevant |
|
703 |
max_new_relevant_facts_per_iter |
|
704 |
(the_default prefers_theory_relevant theory_relevant) |
|
705 |
relevance_override goal hyp_ts concl_t |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
706 |
val (internal_thm_names, formulas) = |
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
707 |
prepare_formulas ctxt full_types hyp_ts concl_t the_axioms |
38023 | 708 |
|
709 |
(* path to unique problem file *) |
|
710 |
val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER" |
|
711 |
else Config.get ctxt dest_dir; |
|
712 |
val the_problem_prefix = Config.get ctxt problem_prefix; |
|
713 |
fun prob_pathname nr = |
|
714 |
let |
|
715 |
val probfile = |
|
716 |
Path.basic ((if overlord then "prob_" ^ name |
|
717 |
else the_problem_prefix ^ serial_string ()) |
|
718 |
^ "_" ^ string_of_int nr) |
|
719 |
in |
|
720 |
if the_dest_dir = "" then File.tmp_path probfile |
|
721 |
else if File.exists (Path.explode the_dest_dir) |
|
722 |
then Path.append (Path.explode the_dest_dir) probfile |
|
723 |
else error ("No such directory: " ^ the_dest_dir ^ ".") |
|
724 |
end; |
|
725 |
||
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38091
diff
changeset
|
726 |
val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec) |
38023 | 727 |
(* write out problem file and call prover *) |
728 |
fun command_line complete probfile = |
|
729 |
let |
|
730 |
val core = File.shell_path command ^ " " ^ arguments complete timeout ^ |
|
731 |
" " ^ File.shell_path probfile |
|
732 |
in |
|
733 |
(if Config.get ctxt measure_runtime then |
|
734 |
"TIMEFORMAT='%3U'; { time " ^ core ^ " ; }" |
|
735 |
else |
|
736 |
"exec " ^ core) ^ " 2>&1" |
|
737 |
end |
|
738 |
fun split_time s = |
|
739 |
let |
|
740 |
val split = String.tokens (fn c => str c = "\n"); |
|
741 |
val (output, t) = s |> split |> split_last |> apfst cat_lines; |
|
742 |
fun as_num f = f >> (fst o read_int); |
|
743 |
val num = as_num (Scan.many1 Symbol.is_ascii_digit); |
|
744 |
val digit = Scan.one Symbol.is_ascii_digit; |
|
745 |
val num3 = as_num (digit ::: digit ::: (digit >> single)); |
|
746 |
val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b); |
|
747 |
val as_time = the_default 0 o Scan.read Symbol.stopper time o explode; |
|
748 |
in (output, as_time t) end; |
|
749 |
fun run_on probfile = |
|
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38091
diff
changeset
|
750 |
case filter (curry (op =) "" o getenv o fst) (exec :: required_execs) of |
38032 | 751 |
(home_var, _) :: _ => |
38023 | 752 |
error ("The environment variable " ^ quote home_var ^ " is not set.") |
38032 | 753 |
| [] => |
754 |
if File.exists command then |
|
755 |
let |
|
756 |
fun do_run complete = |
|
757 |
let |
|
758 |
val command = command_line complete probfile |
|
759 |
val ((output, msecs), res_code) = |
|
760 |
bash_output command |
|
761 |
|>> (if overlord then |
|
762 |
prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n") |
|
763 |
else |
|
764 |
I) |
|
765 |
|>> (if Config.get ctxt measure_runtime then split_time |
|
766 |
else rpair 0) |
|
767 |
val (proof, outcome) = |
|
768 |
extract_proof_and_outcome complete res_code proof_delims |
|
769 |
known_failures output |
|
770 |
in (output, msecs, proof, outcome) end |
|
771 |
val readable_names = debug andalso overlord |
|
772 |
val (pool, conjecture_offset) = |
|
773 |
write_tptp_file thy readable_names explicit_forall full_types |
|
38085
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents:
38084
diff
changeset
|
774 |
explicit_apply probfile formulas |
38032 | 775 |
val conjecture_shape = |
776 |
conjecture_offset + 1 upto conjecture_offset + length hyp_ts + 1 |
|
38040
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents:
38039
diff
changeset
|
777 |
|> map single |
38032 | 778 |
val result = |
779 |
do_run false |
|
780 |
|> (fn (_, msecs0, _, SOME _) => |
|
781 |
do_run true |
|
782 |
|> (fn (output, msecs, proof, outcome) => |
|
783 |
(output, msecs0 + msecs, proof, outcome)) |
|
784 |
| result => result) |
|
785 |
in ((pool, conjecture_shape), result) end |
|
786 |
else |
|
787 |
error ("Bad executable: " ^ Path.implode command ^ ".") |
|
38023 | 788 |
|
789 |
(* If the problem file has not been exported, remove it; otherwise, export |
|
790 |
the proof file too. *) |
|
791 |
fun cleanup probfile = |
|
792 |
if the_dest_dir = "" then try File.rm probfile else NONE |
|
793 |
fun export probfile (_, (output, _, _, _)) = |
|
794 |
if the_dest_dir = "" then |
|
795 |
() |
|
796 |
else |
|
797 |
File.write (Path.explode (Path.implode probfile ^ "_proof")) output |
|
798 |
||
799 |
val ((pool, conjecture_shape), (output, msecs, proof, outcome)) = |
|
800 |
with_path cleanup export run_on (prob_pathname subgoal) |
|
801 |
val (conjecture_shape, internal_thm_names) = |
|
802 |
repair_conjecture_shape_and_theorem_names output conjecture_shape |
|
803 |
internal_thm_names |
|
804 |
||
805 |
val (message, used_thm_names) = |
|
806 |
case outcome of |
|
807 |
NONE => |
|
808 |
proof_text isar_proof |
|
809 |
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape) |
|
810 |
(full_types, minimize_command, proof, internal_thm_names, th, |
|
811 |
subgoal) |
|
812 |
| SOME failure => (string_for_failure failure ^ "\n", []) |
|
813 |
in |
|
814 |
{outcome = outcome, message = message, pool = pool, |
|
815 |
used_thm_names = used_thm_names, atp_run_time_in_msecs = msecs, |
|
816 |
output = output, proof = proof, internal_thm_names = internal_thm_names, |
|
38083
c4b57f68ddb3
remove the "extra_clauses" business introduced in 19a5f1c8a844;
blanchet
parents:
38061
diff
changeset
|
817 |
conjecture_shape = conjecture_shape} |
38023 | 818 |
end |
819 |
||
820 |
fun get_prover_fun thy name = prover_fun name (get_prover thy name) |
|
821 |
||
37584 | 822 |
fun start_prover_thread (params as {verbose, full_types, timeout, ...}) i n |
823 |
relevance_override minimize_command proof_state name = |
|
36379
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
824 |
let |
38023 | 825 |
val thy = Proof.theory_of proof_state |
37584 | 826 |
val birth_time = Time.now () |
827 |
val death_time = Time.+ (birth_time, timeout) |
|
38023 | 828 |
val prover = get_prover_fun thy name |
36379
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
829 |
val {context = ctxt, facts, goal} = Proof.goal proof_state; |
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
830 |
val desc = |
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
831 |
"ATP " ^ quote name ^ " for subgoal " ^ string_of_int i ^ ":\n" ^ |
36392 | 832 |
Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)); |
37584 | 833 |
in |
37585 | 834 |
Async_Manager.launch das_Tool verbose birth_time death_time desc |
37584 | 835 |
(fn () => |
836 |
let |
|
837 |
val problem = |
|
838 |
{subgoal = i, goal = (ctxt, (facts, goal)), |
|
38084
e2aac207d13b
"axiom_clauses" -> "axioms" (these are no longer clauses)
blanchet
parents:
38083
diff
changeset
|
839 |
relevance_override = relevance_override, axioms = NONE} |
37584 | 840 |
in |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38098
diff
changeset
|
841 |
prover params (minimize_command name) problem |> #message |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37926
diff
changeset
|
842 |
handle ERROR message => "Error: " ^ message ^ "\n" |
37584 | 843 |
end) |
844 |
end |
|
28582 | 845 |
|
38044 | 846 |
fun run_sledgehammer {atps = [], ...} _ _ _ _ = error "No ATP is set." |
847 |
| run_sledgehammer (params as {atps, ...}) i relevance_override |
|
848 |
minimize_command state = |
|
849 |
case subgoal_count state of |
|
850 |
0 => priority "No subgoal!" |
|
851 |
| n => |
|
852 |
let |
|
853 |
val _ = kill_atps () |
|
854 |
val _ = priority "Sledgehammering..." |
|
855 |
val _ = app (start_prover_thread params i n relevance_override |
|
856 |
minimize_command state) atps |
|
857 |
in () end |
|
858 |
||
38023 | 859 |
val setup = |
860 |
dest_dir_setup |
|
861 |
#> problem_prefix_setup |
|
862 |
#> measure_runtime_setup |
|
863 |
||
28582 | 864 |
end; |