author | wenzelm |
Fri, 07 Jun 2024 23:53:31 +0200 | |
changeset 80295 | 8a9588ffc133 |
parent 75344 | 647611e6da76 |
permissions | -rw-r--r-- |
55208 | 1 |
(* Title: HOL/TPTP/atp_problem_import.ML |
46324 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
3 |
Copyright 2012 |
|
4 |
||
5 |
Import TPTP problems as Isabelle terms or goals. |
|
6 |
*) |
|
7 |
||
8 |
signature ATP_PROBLEM_IMPORT = |
|
9 |
sig |
|
54434 | 10 |
val read_tptp_file : theory -> (string * term -> 'a) -> string -> |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
11 |
'a list * ('a list * 'a list) * local_theory |
47794
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
blanchet
parents:
47793
diff
changeset
|
12 |
val nitpick_tptp_file : theory -> int -> string -> unit |
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
blanchet
parents:
47793
diff
changeset
|
13 |
val refute_tptp_file : theory -> int -> string -> unit |
64561 | 14 |
val can_tac : Proof.context -> (Proof.context -> tactic) -> term -> bool |
47845 | 15 |
val SOLVE_TIMEOUT : int -> string -> tactic -> tactic |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
16 |
val atp_tac : local_theory -> int -> (string * string) list -> int -> term list -> string -> |
57812
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
17 |
int -> tactic |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
18 |
val smt_solver_tac : string -> local_theory -> int -> tactic |
47845 | 19 |
val make_conj : term list * term list -> term list -> term |
47794
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
blanchet
parents:
47793
diff
changeset
|
20 |
val sledgehammer_tptp_file : theory -> int -> string -> unit |
48083
a4148c95134d
renamed TPTP commands to agree with Sutcliffe's terminology
blanchet
parents:
48082
diff
changeset
|
21 |
val isabelle_tptp_file : theory -> int -> string -> unit |
a4148c95134d
renamed TPTP commands to agree with Sutcliffe's terminology
blanchet
parents:
48082
diff
changeset
|
22 |
val isabelle_hot_tptp_file : theory -> int -> string -> unit |
54472
073f041d83ae
send output of "tptp_translate" to standard output, to simplify Geoff Sutcliffe's life
blanchet
parents:
54434
diff
changeset
|
23 |
val translate_tptp_file : theory -> string -> string -> unit |
46324 | 24 |
end; |
25 |
||
26 |
structure ATP_Problem_Import : ATP_PROBLEM_IMPORT = |
|
27 |
struct |
|
28 |
||
47643
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents:
47565
diff
changeset
|
29 |
open ATP_Util |
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents:
47565
diff
changeset
|
30 |
open ATP_Problem |
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents:
47565
diff
changeset
|
31 |
open ATP_Proof |
54434 | 32 |
open ATP_Problem_Generate |
47643
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents:
47565
diff
changeset
|
33 |
|
47776
024cf0f7fb6d
further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents:
47771
diff
changeset
|
34 |
val debug = false |
024cf0f7fb6d
further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents:
47771
diff
changeset
|
35 |
val overlord = false |
024cf0f7fb6d
further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents:
47771
diff
changeset
|
36 |
|
47643
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents:
47565
diff
changeset
|
37 |
|
47771 | 38 |
(** TPTP parsing **) |
47643
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents:
47565
diff
changeset
|
39 |
|
54434 | 40 |
fun exploded_absolute_path file_name = |
41 |
Path.explode file_name |
|
42 |
|> (fn path => path |> not (Path.is_absolute path) ? Path.append (Path.explode "$PWD")) |
|
43 |
||
47785 | 44 |
fun read_tptp_file thy postproc file_name = |
47644 | 45 |
let |
47718
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents:
47715
diff
changeset
|
46 |
fun has_role role (_, role', _, _) = (role' = role) |
57809 | 47 |
fun get_prop f (name, _, P, _) = P |> f |> close_form |> pair name |> postproc |
48 |
||
54434 | 49 |
val path = exploded_absolute_path file_name |
47714 | 50 |
val ((_, _, problem), thy) = |
57809 | 51 |
TPTP_Interpret.interpret_file true [Path.dir path, Path.explode "$TPTP"] path [] [] thy |
52 |
val (conjs, defs_and_nondefs) = problem |
|
53 |
|> List.partition (has_role TPTP_Syntax.Role_Conjecture) |
|
54 |
||> List.partition (has_role TPTP_Syntax.Role_Definition) |
|
47644 | 55 |
in |
57809 | 56 |
(map (get_prop I) conjs, |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58843
diff
changeset
|
57 |
apply2 (map (get_prop Logic.varify_global)) defs_and_nondefs, |
57812
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
58 |
Named_Target.theory_init thy) |
47557
32f35b3d9e42
started integrating Nik's parser into TPTP command-line tools
blanchet
parents:
46325
diff
changeset
|
59 |
end |
46325 | 60 |
|
47771 | 61 |
|
57547
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
blanchet
parents:
57268
diff
changeset
|
62 |
(** Nitpick **) |
46324 | 63 |
|
74402 | 64 |
fun aptrueprop f ((t0 as \<^Const_>\<open>Trueprop\<close>) $ t1) = t0 $ f t1 |
47718
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents:
47715
diff
changeset
|
65 |
| aptrueprop f t = f t |
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents:
47715
diff
changeset
|
66 |
|
74402 | 67 |
fun is_legitimate_tptp_def \<^Const_>\<open>Trueprop for t\<close> = is_legitimate_tptp_def t |
68 |
| is_legitimate_tptp_def \<^Const_>\<open>HOL.eq _ for t u\<close> = |
|
57547
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
blanchet
parents:
57268
diff
changeset
|
69 |
(is_Const t orelse is_Free t) andalso not (exists_subterm (curry (op =) t) u) |
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
blanchet
parents:
57268
diff
changeset
|
70 |
| is_legitimate_tptp_def _ = false |
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
blanchet
parents:
57268
diff
changeset
|
71 |
|
47794
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
blanchet
parents:
47793
diff
changeset
|
72 |
fun nitpick_tptp_file thy timeout file_name = |
46325 | 73 |
let |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
74 |
val (conjs, (defs, nondefs), lthy) = read_tptp_file thy snd file_name |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
75 |
val thy = Proof_Context.theory_of lthy |
57809 | 76 |
val (defs, pseudo_defs) = defs |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
77 |
|> map (ATP_Util.abs_extensionalize_term lthy #> aptrueprop (hol_open_form I)) |
57809 | 78 |
|> List.partition (is_legitimate_tptp_def o perhaps (try HOLogic.dest_Trueprop) |
79 |
o ATP_Util.unextensionalize_def) |
|
47991
3eb598b044ad
make Nitpick's handling of definitions more robust in the face of formulas that don't have the expected format (needed for soundness, cf. RNG100+1)
blanchet
parents:
47957
diff
changeset
|
80 |
val nondefs = pseudo_defs @ nondefs |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
81 |
val state = Proof.init lthy |
46325 | 82 |
val params = |
61569
947ce60a06e1
eliminated Nitpick's pedantic support for 'emdash'
blanchet
parents:
61310
diff
changeset
|
83 |
[("card", "1-100"), |
46325 | 84 |
("box", "false"), |
85 |
("max_threads", "1"), |
|
47791 | 86 |
("batch_size", "5"), |
47718
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents:
47715
diff
changeset
|
87 |
("falsify", if null conjs then "false" else "true"), |
46325 | 88 |
("verbose", "true"), |
47776
024cf0f7fb6d
further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents:
47771
diff
changeset
|
89 |
("debug", if debug then "true" else "false"), |
024cf0f7fb6d
further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents:
47771
diff
changeset
|
90 |
("overlord", if overlord then "true" else "false"), |
46325 | 91 |
("show_consts", "true"), |
47755 | 92 |
("format", "1"), |
46325 | 93 |
("max_potential", "0"), |
47718
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents:
47715
diff
changeset
|
94 |
("timeout", string_of_int timeout), |
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents:
47715
diff
changeset
|
95 |
("tac_timeout", string_of_int ((timeout + 49) div 50))] |
55199 | 96 |
|> Nitpick_Commands.default_params thy |
46325 | 97 |
val i = 1 |
98 |
val n = 1 |
|
99 |
val step = 0 |
|
100 |
val subst = [] |
|
101 |
in |
|
57809 | 102 |
Nitpick.pick_nits_in_term state params Nitpick.TPTP i n step subst defs nondefs |
69597 | 103 |
(case conjs of conj :: _ => conj | [] => \<^prop>\<open>True\<close>); |
47718
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents:
47715
diff
changeset
|
104 |
() |
46325 | 105 |
end |
46324 | 106 |
|
107 |
||
108 |
(** Refute **) |
|
109 |
||
47794
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
blanchet
parents:
47793
diff
changeset
|
110 |
fun refute_tptp_file thy timeout file_name = |
46325 | 111 |
let |
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51717
diff
changeset
|
112 |
fun print_szs_of_outcome falsify s = |
47766
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
113 |
"% SZS status " ^ |
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
114 |
(if s = "genuine" then |
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
115 |
if falsify then "CounterSatisfiable" else "Satisfiable" |
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
116 |
else |
61310 | 117 |
"GaveUp") |
58843 | 118 |
|> writeln |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
119 |
val (conjs, (defs, nondefs), lthy) = read_tptp_file thy snd file_name |
46325 | 120 |
val params = |
47714 | 121 |
[("maxtime", string_of_int timeout), |
122 |
("maxvars", "100000")] |
|
46325 | 123 |
in |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
124 |
Refute.refute_term lthy params (defs @ nondefs) |
69597 | 125 |
(case conjs of conj :: _ => conj | [] => \<^prop>\<open>True\<close>) |
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51717
diff
changeset
|
126 |
|> print_szs_of_outcome (not (null conjs)) |
46325 | 127 |
end |
46324 | 128 |
|
129 |
||
47766
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
130 |
(** Sledgehammer and Isabelle (combination of provers) **) |
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
131 |
|
64561 | 132 |
fun can_tac ctxt tactic conj = |
74530
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
wenzelm
parents:
74526
diff
changeset
|
133 |
can (Goal.prove_internal ctxt [] (Thm.cterm_of ctxt conj)) (fn [] => tactic ctxt) |
47771 | 134 |
|
47766
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
135 |
fun SOLVE_TIMEOUT seconds name tac st = |
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
136 |
let |
58843 | 137 |
val _ = writeln ("running " ^ name ^ " for " ^ string_of_int seconds ^ " s") |
47766
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
138 |
val result = |
62519 | 139 |
Timeout.apply (Time.fromSeconds seconds) (fn () => SINGLE (SOLVE tac) st) () |
64561 | 140 |
handle Timeout.TIMEOUT _ => NONE | ERROR _ => NONE |
47766
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
141 |
in |
57809 | 142 |
(case result of |
58843 | 143 |
NONE => (writeln ("FAILURE: " ^ name); Seq.empty) |
144 |
| SOME st' => (writeln ("SUCCESS: " ^ name); Seq.single st')) |
|
47766
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
145 |
end |
46324 | 146 |
|
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
147 |
fun nitpick_finite_oracle_tac lthy timeout i th = |
47812 | 148 |
let |
69597 | 149 |
fun is_safe (Type (\<^type_name>\<open>fun\<close>, Ts)) = forall is_safe Ts |
150 |
| is_safe \<^typ>\<open>prop\<close> = true |
|
151 |
| is_safe \<^typ>\<open>bool\<close> = true |
|
47812 | 152 |
| is_safe _ = false |
57809 | 153 |
|
47812 | 154 |
val conj = Thm.term_of (Thm.cprem_of th i) |
155 |
in |
|
156 |
if exists_type (not o is_safe) conj then |
|
157 |
Seq.empty |
|
158 |
else |
|
159 |
let |
|
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
160 |
val thy = Proof_Context.theory_of lthy |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
161 |
val state = Proof.init lthy |
47812 | 162 |
val params = |
163 |
[("box", "false"), |
|
164 |
("max_threads", "1"), |
|
165 |
("verbose", "true"), |
|
166 |
("debug", if debug then "true" else "false"), |
|
167 |
("overlord", if overlord then "true" else "false"), |
|
168 |
("max_potential", "0"), |
|
169 |
("timeout", string_of_int timeout)] |
|
55199 | 170 |
|> Nitpick_Commands.default_params thy |
47812 | 171 |
val i = 1 |
172 |
val n = 1 |
|
173 |
val step = 0 |
|
174 |
val subst = [] |
|
175 |
val (outcome, _) = |
|
57809 | 176 |
Nitpick.pick_nits_in_term state params Nitpick.Normal i n step subst [] [] conj |
177 |
in |
|
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
178 |
if outcome = "none" then ALLGOALS (Skip_Proof.cheat_tac lthy) th else Seq.empty |
57809 | 179 |
end |
47812 | 180 |
end |
181 |
||
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
182 |
fun atp_tac lthy completeness override_params timeout assms prover = |
47946
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47845
diff
changeset
|
183 |
let |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
184 |
val thy = Proof_Context.theory_of lthy |
57812
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
185 |
val assm_ths0 = map (Skip_Proof.make_thm thy) assms |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
186 |
val ((assm_name, _), lthy) = lthy |
64445 | 187 |
|> Config.put Sledgehammer_Prover_ATP.atp_completish (if completeness > 0 then 3 else 0) |
69597 | 188 |
|> Local_Theory.note ((\<^binding>\<open>thms\<close>, []), assm_ths0) |
57812
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
189 |
|
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
75344
diff
changeset
|
190 |
fun ref_of th = (Facts.named (Thm_Name.short (Thm.derivation_name th)), []) |
57812
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
191 |
val ref_of_assms = (Facts.named assm_name, []) |
47946
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47845
diff
changeset
|
192 |
in |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
193 |
Sledgehammer_Tactics.sledgehammer_as_oracle_tac lthy |
57809 | 194 |
([("debug", if debug then "true" else "false"), |
195 |
("overlord", if overlord then "true" else "false"), |
|
196 |
("provers", prover), |
|
197 |
("timeout", string_of_int timeout)] @ |
|
198 |
(if completeness > 0 then |
|
199 |
[("type_enc", if completeness = 1 then "mono_native" else "poly_tags")] |
|
200 |
else |
|
201 |
[]) @ |
|
202 |
override_params) |
|
57812
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
203 |
{add = ref_of_assms :: map ref_of [ext, @{thm someI}], del = [], only = true} [] |
47946
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47845
diff
changeset
|
204 |
end |
47765 | 205 |
|
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
206 |
fun sledgehammer_tac demo lthy timeout assms i = |
47766
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
207 |
let |
47946
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47845
diff
changeset
|
208 |
val frac = if demo then 16 else 12 |
48143 | 209 |
fun slice mult completeness prover = |
47946
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47845
diff
changeset
|
210 |
SOLVE_TIMEOUT (mult * timeout div frac) |
57809 | 211 |
(prover ^ (if completeness > 0 then "(" ^ string_of_int completeness ^ ")" else "")) |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
212 |
(atp_tac lthy completeness [] (mult * timeout div frac) assms prover i) |
47766
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
213 |
in |
57154 | 214 |
slice 2 0 ATP_Proof.spassN |
215 |
ORELSE slice 2 0 ATP_Proof.vampireN |
|
216 |
ORELSE slice 2 0 ATP_Proof.eN |
|
217 |
ORELSE slice 1 1 ATP_Proof.spassN |
|
218 |
ORELSE slice 1 2 ATP_Proof.eN |
|
219 |
ORELSE slice 1 1 ATP_Proof.vampireN |
|
220 |
ORELSE slice 1 2 ATP_Proof.vampireN |
|
47946
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47845
diff
changeset
|
221 |
ORELSE |
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47845
diff
changeset
|
222 |
(if demo then |
57154 | 223 |
slice 2 0 ATP_Proof.satallaxN |
224 |
ORELSE slice 2 0 ATP_Proof.leo2N |
|
47946
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47845
diff
changeset
|
225 |
else |
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47845
diff
changeset
|
226 |
no_tac) |
47766
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
227 |
end |
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
228 |
|
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
229 |
fun smt_solver_tac solver lthy = |
47832 | 230 |
let |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
231 |
val lthy = lthy |> Context.proof_map (SMT_Config.select_solver solver) |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
232 |
in SMT_Solver.smt_tac lthy [] end |
47832 | 233 |
|
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
234 |
fun auto_etc_tac lthy timeout assms i = |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
235 |
SOLVE_TIMEOUT (timeout div 20) "nitpick" (nitpick_finite_oracle_tac lthy (timeout div 20) i) |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
236 |
ORELSE SOLVE_TIMEOUT (timeout div 10) "simp" (asm_full_simp_tac lthy i) |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
237 |
ORELSE SOLVE_TIMEOUT (timeout div 10) "blast" (blast_tac lthy i) |
47794
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
blanchet
parents:
47793
diff
changeset
|
238 |
ORELSE SOLVE_TIMEOUT (timeout div 5) "auto+spass" |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
239 |
(auto_tac lthy THEN ALLGOALS (atp_tac lthy 0 [] (timeout div 5) assms ATP_Proof.spassN)) |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
240 |
ORELSE SOLVE_TIMEOUT (timeout div 10) "fast" (fast_tac lthy i) |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
241 |
ORELSE SOLVE_TIMEOUT (timeout div 20) "z3" (smt_solver_tac "z3" lthy i) |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
242 |
ORELSE SOLVE_TIMEOUT (timeout div 20) "cvc4" (smt_solver_tac "cvc4" lthy i) |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
243 |
ORELSE SOLVE_TIMEOUT (timeout div 20) "best" (best_tac lthy i) |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
244 |
ORELSE SOLVE_TIMEOUT (timeout div 10) "force" (force_tac lthy i) |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
245 |
ORELSE SOLVE_TIMEOUT (timeout div 10) "meson" (Meson.meson_tac lthy [] i) |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
246 |
ORELSE SOLVE_TIMEOUT (timeout div 10) "fastforce" (fast_force_tac lthy i) |
47771 | 247 |
|
47776
024cf0f7fb6d
further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents:
47771
diff
changeset
|
248 |
fun make_conj (defs, nondefs) conjs = |
69597 | 249 |
Logic.list_implies (rev defs @ rev nondefs, case conjs of conj :: _ => conj | [] => \<^prop>\<open>False\<close>) |
47771 | 250 |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51717
diff
changeset
|
251 |
fun print_szs_of_success conjs success = |
58843 | 252 |
writeln ("% SZS status " ^ |
57809 | 253 |
(if success then |
254 |
if null conjs then "Unsatisfiable" else "Theorem" |
|
255 |
else |
|
61310 | 256 |
"GaveUp")) |
47765 | 257 |
|
47794
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
blanchet
parents:
47793
diff
changeset
|
258 |
fun sledgehammer_tptp_file thy timeout file_name = |
47771 | 259 |
let |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
260 |
val (conjs, assms, lthy) = read_tptp_file thy snd file_name |
57812
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
261 |
val conj = make_conj ([], []) conjs |
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
262 |
val assms = op @ assms |
47771 | 263 |
in |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
264 |
can_tac lthy (fn lthy => sledgehammer_tac true lthy timeout assms 1) conj |
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51717
diff
changeset
|
265 |
|> print_szs_of_success conjs |
47771 | 266 |
end |
47766
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents:
47765
diff
changeset
|
267 |
|
48083
a4148c95134d
renamed TPTP commands to agree with Sutcliffe's terminology
blanchet
parents:
48082
diff
changeset
|
268 |
fun generic_isabelle_tptp_file demo thy timeout file_name = |
47771 | 269 |
let |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
270 |
val (conjs, assms, lthy) = read_tptp_file thy snd file_name |
57812
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
271 |
val conj = make_conj ([], []) conjs |
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
272 |
val full_conj = make_conj assms conjs |
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
57809
diff
changeset
|
273 |
val assms = op @ assms |
48143 | 274 |
val (last_hope_atp, last_hope_completeness) = |
57154 | 275 |
if demo then (ATP_Proof.satallaxN, 0) else (ATP_Proof.vampireN, 2) |
47771 | 276 |
in |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
277 |
(can_tac lthy (fn lthy => auto_etc_tac lthy (timeout div 2) assms 1) full_conj orelse |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
278 |
can_tac lthy (fn lthy => sledgehammer_tac demo lthy (timeout div 2) assms 1) conj orelse |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
279 |
can_tac lthy (fn lthy => SOLVE_TIMEOUT timeout (last_hope_atp ^ "(*)") |
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
280 |
(atp_tac lthy last_hope_completeness [] timeout assms last_hope_atp 1)) full_conj) |
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51717
diff
changeset
|
281 |
|> print_szs_of_success conjs |
47771 | 282 |
end |
46324 | 283 |
|
48083
a4148c95134d
renamed TPTP commands to agree with Sutcliffe's terminology
blanchet
parents:
48082
diff
changeset
|
284 |
val isabelle_tptp_file = generic_isabelle_tptp_file false |
a4148c95134d
renamed TPTP commands to agree with Sutcliffe's terminology
blanchet
parents:
48082
diff
changeset
|
285 |
val isabelle_hot_tptp_file = generic_isabelle_tptp_file true |
47832 | 286 |
|
287 |
||
46324 | 288 |
(** Translator between TPTP(-like) file formats **) |
289 |
||
54472
073f041d83ae
send output of "tptp_translate" to standard output, to simplify Geoff Sutcliffe's life
blanchet
parents:
54434
diff
changeset
|
290 |
fun translate_tptp_file thy format_str file_name = |
54434 | 291 |
let |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
292 |
val (conjs, (defs, nondefs), lthy) = read_tptp_file thy I file_name |
54434 | 293 |
val conj = make_conj ([], []) (map snd conjs) |
294 |
||
295 |
val (format, type_enc, lam_trans) = |
|
296 |
(case format_str of |
|
297 |
"FOF" => (FOF, "mono_guards??", liftingN) |
|
74902 | 298 |
| "TF0" => (TFF (Monomorphic, Without_FOOL), "mono_native", liftingN) |
299 |
| "TH0" => (THF (Monomorphic, {with_ite = false, with_let = false}, THF_Without_Choice), |
|
300 |
"mono_native_higher", keep_lamsN) |
|
54434 | 301 |
| "DFG" => (DFG Monomorphic, "mono_native", liftingN) |
302 |
| _ => error ("Unknown format " ^ quote format_str ^ |
|
54547
c999e2533487
renamed TFF0/THF0 to three-letter acronyms, in keeping with new TPTP policy
blanchet
parents:
54472
diff
changeset
|
303 |
" (expected \"FOF\", \"TF0\", \"TH0\", or \"DFG\")")) |
61860
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents:
61569
diff
changeset
|
304 |
val generate_info = false |
54434 | 305 |
val uncurried_aliases = false |
306 |
val readable_names = true |
|
307 |
val presimp = true |
|
57809 | 308 |
val facts = |
309 |
map (apfst (rpair (Global, Non_Rec_Def))) defs @ |
|
54434 | 310 |
map (apfst (rpair (Global, General))) nondefs |
57268 | 311 |
val (atp_problem, _, _, _) = |
72001
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
blanchet
parents:
69597
diff
changeset
|
312 |
generate_atp_problem lthy generate_info format Hypothesis (type_enc_of_string Strict type_enc) |
61860
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents:
61569
diff
changeset
|
313 |
Translator |
54434 | 314 |
lam_trans uncurried_aliases readable_names presimp [] conj facts |
315 |
||
75344 | 316 |
val lines = lines_of_atp_problem format (K []) atp_problem |
54434 | 317 |
in |
54472
073f041d83ae
send output of "tptp_translate" to standard output, to simplify Geoff Sutcliffe's life
blanchet
parents:
54434
diff
changeset
|
318 |
List.app Output.physical_stdout lines |
54434 | 319 |
end |
46324 | 320 |
|
321 |
end; |