| author | blanchet | 
| Thu, 02 Aug 2012 10:10:29 +0200 | |
| changeset 48653 | 6ac7fd9b3a54 | 
| parent 48143 | 0186df5074c8 | 
| child 48829 | 6ed588c4f963 | 
| permissions | -rw-r--r-- | 
| 46324 | 1  | 
(* Title: HOL/TPTP/atp_problem_import.ML  | 
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  | 
|
| 47845 | 10  | 
val read_tptp_file :  | 
11  | 
theory -> (term -> term) -> string  | 
|
12  | 
-> term list * (term list * term list) * Proof.context  | 
|
| 
47794
 
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 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
 | 
14  | 
val refute_tptp_file : theory -> int -> string -> unit  | 
| 47845 | 15  | 
val can_tac : Proof.context -> tactic -> term -> bool  | 
16  | 
val SOLVE_TIMEOUT : int -> string -> tactic -> tactic  | 
|
17  | 
val atp_tac :  | 
|
| 
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
 | 
18  | 
Proof.context -> int -> (string * string) list -> int -> string -> int  | 
| 
 
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
 | 
19  | 
-> tactic  | 
| 47845 | 20  | 
val smt_solver_tac : string -> Proof.context -> int -> tactic  | 
21  | 
val freeze_problem_consts : theory -> term -> term  | 
|
22  | 
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
 | 
23  | 
val sledgehammer_tptp_file : theory -> int -> string -> unit  | 
| 
48083
 
a4148c95134d
renamed TPTP commands to agree with Sutcliffe's terminology
 
blanchet 
parents: 
48082 
diff
changeset
 | 
24  | 
val isabelle_tptp_file : theory -> int -> string -> unit  | 
| 
 
a4148c95134d
renamed TPTP commands to agree with Sutcliffe's terminology
 
blanchet 
parents: 
48082 
diff
changeset
 | 
25  | 
val isabelle_hot_tptp_file : theory -> int -> string -> unit  | 
| 46325 | 26  | 
val translate_tptp_file : string -> string -> string -> unit  | 
| 46324 | 27  | 
end;  | 
28  | 
||
29  | 
structure ATP_Problem_Import : ATP_PROBLEM_IMPORT =  | 
|
30  | 
struct  | 
|
31  | 
||
| 
47643
 
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
 
blanchet 
parents: 
47565 
diff
changeset
 | 
32  | 
open ATP_Util  | 
| 
 
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
 
blanchet 
parents: 
47565 
diff
changeset
 | 
33  | 
open ATP_Problem  | 
| 
 
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
 
blanchet 
parents: 
47565 
diff
changeset
 | 
34  | 
open ATP_Proof  | 
| 
 
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
 
blanchet 
parents: 
47565 
diff
changeset
 | 
35  | 
|
| 
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
 | 
36  | 
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
 | 
37  | 
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
 | 
38  | 
|
| 
47643
 
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
 
blanchet 
parents: 
47565 
diff
changeset
 | 
39  | 
|
| 47771 | 40  | 
(** TPTP parsing **)  | 
| 
47643
 
e33c2be488fe
reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
 
blanchet 
parents: 
47565 
diff
changeset
 | 
41  | 
|
| 
47565
 
762eb0dacaa6
remove old TPTP CNF/FOF parser; always use Nik's new parser
 
blanchet 
parents: 
47562 
diff
changeset
 | 
42  | 
(* cf. "close_form" in "refute.ML" *)  | 
| 
 
762eb0dacaa6
remove old TPTP CNF/FOF parser; always use Nik's new parser
 
blanchet 
parents: 
47562 
diff
changeset
 | 
43  | 
fun close_form t =  | 
| 
 
762eb0dacaa6
remove old TPTP CNF/FOF parser; always use Nik's new parser
 
blanchet 
parents: 
47562 
diff
changeset
 | 
44  | 
fold (fn ((s, i), T) => fn t' =>  | 
| 
 
762eb0dacaa6
remove old TPTP CNF/FOF parser; always use Nik's new parser
 
blanchet 
parents: 
47562 
diff
changeset
 | 
45  | 
Logic.all_const T $ Abs (s, T, abstract_over (Var ((s, i), T), t')))  | 
| 
 
762eb0dacaa6
remove old TPTP CNF/FOF parser; always use Nik's new parser
 
blanchet 
parents: 
47562 
diff
changeset
 | 
46  | 
(Term.add_vars t []) t  | 
| 46324 | 47  | 
|
| 47785 | 48  | 
fun read_tptp_file thy postproc file_name =  | 
| 47644 | 49  | 
let  | 
| 
47718
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
50  | 
fun has_role role (_, role', _, _) = (role' = role)  | 
| 47785 | 51  | 
fun get_prop (_, _, P, _) =  | 
52  | 
P |> Logic.varify_global |> close_form |> postproc  | 
|
| 47644 | 53  | 
val path =  | 
54  | 
Path.explode file_name  | 
|
55  | 
|> (fn path =>  | 
|
56  | 
path |> not (Path.is_absolute path)  | 
|
57  | 
? Path.append (Path.explode "$PWD"))  | 
|
| 47714 | 58  | 
val ((_, _, problem), thy) =  | 
59  | 
TPTP_Interpret.interpret_file true (Path.explode "$TPTP") path [] [] thy  | 
|
| 
47718
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
60  | 
val (conjs, defs_and_nondefs) =  | 
| 
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
61  | 
problem |> List.partition (has_role TPTP_Syntax.Role_Conjecture)  | 
| 
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
62  | 
||> List.partition (has_role TPTP_Syntax.Role_Definition)  | 
| 47644 | 63  | 
in  | 
| 
47718
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
64  | 
(map get_prop conjs, pairself (map get_prop) defs_and_nondefs,  | 
| 47771 | 65  | 
thy |> Theory.checkpoint |> Proof_Context.init_global)  | 
| 
47557
 
32f35b3d9e42
started integrating Nik's parser into TPTP command-line tools
 
blanchet 
parents: 
46325 
diff
changeset
 | 
66  | 
end  | 
| 46325 | 67  | 
|
| 47771 | 68  | 
|
| 46324 | 69  | 
(** Nitpick (alias Nitrox) **)  | 
70  | 
||
| 
47718
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
71  | 
fun aptrueprop f ((t0 as @{const Trueprop}) $ t1) = t0 $ f t1
 | 
| 
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
72  | 
| aptrueprop f t = f t  | 
| 
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
73  | 
|
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
74  | 
fun nitpick_tptp_file thy timeout file_name =  | 
| 46325 | 75  | 
let  | 
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
76  | 
val (conjs, (defs, nondefs), ctxt) = read_tptp_file thy I file_name  | 
| 47771 | 77  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
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
 | 
78  | 
val (defs, pseudo_defs) =  | 
| 
 
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
 | 
79  | 
defs |> map (ATP_Util.abs_extensionalize_term ctxt  | 
| 
 
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  | 
#> aptrueprop (open_form I))  | 
| 
 
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
 | 
81  | 
|> List.partition (ATP_Util.is_legitimate_tptp_def  | 
| 
 
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
 | 
82  | 
o perhaps (try HOLogic.dest_Trueprop)  | 
| 
 
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
 | 
83  | 
o ATP_Util.unextensionalize_def)  | 
| 
 
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
 | 
84  | 
val nondefs = pseudo_defs @ nondefs  | 
| 47714 | 85  | 
val state = Proof.init ctxt  | 
| 46325 | 86  | 
val params =  | 
| 
47557
 
32f35b3d9e42
started integrating Nik's parser into TPTP command-line tools
 
blanchet 
parents: 
46325 
diff
changeset
 | 
87  | 
      [("card", "1\<emdash>100"),
 | 
| 46325 | 88  | 
       ("box", "false"),
 | 
89  | 
       ("max_threads", "1"),
 | 
|
| 47791 | 90  | 
       ("batch_size", "5"),
 | 
| 
47718
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
91  | 
       ("falsify", if null conjs then "false" else "true"),
 | 
| 46325 | 92  | 
       ("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
 | 
93  | 
       ("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
 | 
94  | 
       ("overlord", if overlord then "true" else "false"),
 | 
| 46325 | 95  | 
       ("show_consts", "true"),
 | 
| 47755 | 96  | 
       ("format", "1"),
 | 
| 46325 | 97  | 
       ("max_potential", "0"),
 | 
| 
47718
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
98  | 
       ("timeout", string_of_int timeout),
 | 
| 
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
99  | 
       ("tac_timeout", string_of_int ((timeout + 49) div 50))]
 | 
| 47714 | 100  | 
|> Nitpick_Isar.default_params thy  | 
| 46325 | 101  | 
val i = 1  | 
102  | 
val n = 1  | 
|
103  | 
val step = 0  | 
|
104  | 
val subst = []  | 
|
105  | 
in  | 
|
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47714 
diff
changeset
 | 
106  | 
Nitpick.pick_nits_in_term state params Nitpick.TPTP i n step subst  | 
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
107  | 
        defs nondefs (case conjs of conj :: _ => conj | [] => @{prop True});
 | 
| 
47718
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
108  | 
()  | 
| 46325 | 109  | 
end  | 
| 46324 | 110  | 
|
111  | 
||
112  | 
(** Refute **)  | 
|
113  | 
||
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
114  | 
fun refute_tptp_file thy timeout file_name =  | 
| 46325 | 115  | 
let  | 
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
116  | 
fun print_szs_from_outcome falsify s =  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
117  | 
"% SZS status " ^  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
118  | 
(if s = "genuine" then  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
119  | 
if falsify then "CounterSatisfiable" else "Satisfiable"  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
120  | 
else  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
121  | 
"Unknown")  | 
| 47788 | 122  | 
|> Output.urgent_message  | 
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
123  | 
val (conjs, (defs, nondefs), ctxt) = read_tptp_file thy I file_name  | 
| 46325 | 124  | 
val params =  | 
| 47714 | 125  | 
      [("maxtime", string_of_int timeout),
 | 
126  | 
       ("maxvars", "100000")]
 | 
|
| 46325 | 127  | 
in  | 
| 
47718
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
128  | 
Refute.refute_term ctxt params (defs @ nondefs)  | 
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
129  | 
        (case conjs of conj :: _ => conj | [] => @{prop True})
 | 
| 
47718
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
130  | 
|> print_szs_from_outcome (not (null conjs))  | 
| 46325 | 131  | 
end  | 
| 46324 | 132  | 
|
133  | 
||
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
134  | 
(** Sledgehammer and Isabelle (combination of provers) **)  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
135  | 
|
| 47785 | 136  | 
fun can_tac ctxt tactic conj = can (Goal.prove ctxt [] [] conj) (K tactic)  | 
| 47771 | 137  | 
|
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
138  | 
fun SOLVE_TIMEOUT seconds name tac st =  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
139  | 
let  | 
| 47788 | 140  | 
    val _ = Output.urgent_message ("running " ^ name ^ " for " ^
 | 
141  | 
string_of_int seconds ^ " s")  | 
|
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
142  | 
val result =  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
143  | 
TimeLimit.timeLimit (Time.fromSeconds seconds)  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
144  | 
(fn () => SINGLE (SOLVE tac) st) ()  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
145  | 
handle TimeLimit.TimeOut => NONE  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
146  | 
| ERROR _ => NONE  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
147  | 
in  | 
| 47785 | 148  | 
case result of  | 
| 47788 | 149  | 
      NONE => (Output.urgent_message ("FAILURE: " ^ name); Seq.empty)
 | 
150  | 
    | SOME st' => (Output.urgent_message ("SUCCESS: " ^ name); Seq.single st')
 | 
|
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
151  | 
end  | 
| 46324 | 152  | 
|
| 47812 | 153  | 
fun nitpick_finite_oracle_tac ctxt timeout i th =  | 
154  | 
let  | 
|
155  | 
    fun is_safe (Type (@{type_name fun}, Ts)) = forall is_safe Ts
 | 
|
156  | 
      | is_safe @{typ prop} = true
 | 
|
157  | 
      | is_safe @{typ bool} = true
 | 
|
158  | 
| is_safe _ = false  | 
|
159  | 
val conj = Thm.term_of (Thm.cprem_of th i)  | 
|
160  | 
in  | 
|
161  | 
if exists_type (not o is_safe) conj then  | 
|
162  | 
Seq.empty  | 
|
163  | 
else  | 
|
164  | 
let  | 
|
165  | 
val thy = Proof_Context.theory_of ctxt  | 
|
166  | 
val state = Proof.init ctxt  | 
|
167  | 
val params =  | 
|
168  | 
          [("box", "false"),
 | 
|
169  | 
           ("max_threads", "1"),
 | 
|
170  | 
           ("verbose", "true"),
 | 
|
171  | 
           ("debug", if debug then "true" else "false"),
 | 
|
172  | 
           ("overlord", if overlord then "true" else "false"),
 | 
|
173  | 
           ("max_potential", "0"),
 | 
|
174  | 
           ("timeout", string_of_int timeout)]
 | 
|
175  | 
|> Nitpick_Isar.default_params thy  | 
|
176  | 
val i = 1  | 
|
177  | 
val n = 1  | 
|
178  | 
val step = 0  | 
|
179  | 
val subst = []  | 
|
180  | 
val (outcome, _) =  | 
|
181  | 
Nitpick.pick_nits_in_term state params Nitpick.Normal i n step subst  | 
|
182  | 
[] [] conj  | 
|
183  | 
in if outcome = "none" then Skip_Proof.cheat_tac thy th else Seq.empty end  | 
|
184  | 
end  | 
|
185  | 
||
| 48143 | 186  | 
fun atp_tac ctxt completeness override_params timeout 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
 | 
187  | 
let  | 
| 
 
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
 | 
188  | 
val ctxt =  | 
| 48143 | 189  | 
ctxt |> Config.put Sledgehammer_Provers.completish (completeness > 0)  | 
| 
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
 | 
190  | 
in  | 
| 
 
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
 | 
191  | 
Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt  | 
| 
 
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  | 
        ([("debug", if debug then "true" else "false"),
 | 
| 
 
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
 | 
193  | 
          ("overlord", if overlord then "true" else "false"),
 | 
| 
 
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
 | 
194  | 
          ("provers", prover),
 | 
| 
 
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
 | 
195  | 
          ("timeout", string_of_int timeout)] @
 | 
| 48143 | 196  | 
(if completeness > 0 then  | 
| 
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
 | 
197  | 
            [("type_enc",
 | 
| 48143 | 198  | 
if completeness = 1 then "mono_native" else "poly_guards??"),  | 
| 
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
 | 
199  | 
             ("slicing", "false")]
 | 
| 
 
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
 | 
200  | 
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
 | 
201  | 
[]) @  | 
| 
 
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
 | 
202  | 
override_params)  | 
| 
 
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
 | 
203  | 
        {add = [(Facts.named (Thm.derivation_name ext), [])],
 | 
| 
 
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  | 
del = [], only = true}  | 
| 
 
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
 | 
205  | 
end  | 
| 47765 | 206  | 
|
| 47832 | 207  | 
fun sledgehammer_tac demo ctxt timeout i =  | 
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
208  | 
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
 | 
209  | 
val frac = if demo then 16 else 12  | 
| 48143 | 210  | 
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
 | 
211  | 
SOLVE_TIMEOUT (mult * timeout div frac)  | 
| 
 
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
 | 
212  | 
(prover ^  | 
| 48143 | 213  | 
           (if completeness > 0 then "(" ^ string_of_int completeness ^ ")"
 | 
| 
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
 | 
214  | 
else ""))  | 
| 48143 | 215  | 
(atp_tac ctxt completeness [] (mult * timeout div frac) prover i)  | 
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
216  | 
in  | 
| 
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
 | 
217  | 
slice 2 0 ATP_Systems.spassN  | 
| 
 
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
 | 
218  | 
ORELSE slice 2 0 ATP_Systems.vampireN  | 
| 
 
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
 | 
219  | 
ORELSE slice 2 0 ATP_Systems.eN  | 
| 
 
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
 | 
220  | 
ORELSE slice 2 0 ATP_Systems.z3_tptpN  | 
| 
 
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 slice 1 1 ATP_Systems.spassN  | 
| 
 
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  | 
ORELSE slice 1 2 ATP_Systems.eN  | 
| 
 
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
 | 
223  | 
ORELSE slice 1 1 ATP_Systems.vampireN  | 
| 
 
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
 | 
224  | 
ORELSE slice 1 2 ATP_Systems.vampireN  | 
| 
 
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  | 
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
 | 
226  | 
(if demo then  | 
| 
 
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
 | 
227  | 
slice 2 0 ATP_Systems.satallaxN  | 
| 
 
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
 | 
228  | 
ORELSE slice 2 0 ATP_Systems.leo2N  | 
| 
 
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
 | 
229  | 
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
 | 
230  | 
no_tac)  | 
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
231  | 
end  | 
| 
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
232  | 
|
| 47832 | 233  | 
fun smt_solver_tac solver ctxt =  | 
234  | 
let  | 
|
235  | 
val ctxt = ctxt |> Context.proof_map (SMT_Config.select_solver solver)  | 
|
236  | 
in SMT_Solver.smt_tac ctxt [] end  | 
|
237  | 
||
| 47785 | 238  | 
fun auto_etc_tac ctxt timeout i =  | 
| 47812 | 239  | 
SOLVE_TIMEOUT (timeout div 20) "nitpick"  | 
240  | 
(nitpick_finite_oracle_tac ctxt (timeout div 20) i)  | 
|
| 
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
 | 
241  | 
ORELSE SOLVE_TIMEOUT (timeout div 10) "simp"  | 
| 47812 | 242  | 
(asm_full_simp_tac (simpset_of ctxt) i)  | 
| 47788 | 243  | 
ORELSE SOLVE_TIMEOUT (timeout div 10) "blast" (blast_tac ctxt 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
 | 
244  | 
ORELSE SOLVE_TIMEOUT (timeout div 5) "auto+spass"  | 
| 47785 | 245  | 
(auto_tac ctxt  | 
| 
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
 | 
246  | 
THEN ALLGOALS (atp_tac ctxt 0 [] (timeout div 5) ATP_Systems.spassN))  | 
| 
 
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
 | 
247  | 
ORELSE SOLVE_TIMEOUT (timeout div 10) "fast" (fast_tac ctxt i)  | 
| 
 
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
 | 
248  | 
ORELSE SOLVE_TIMEOUT (timeout div 20) "z3" (smt_solver_tac "z3" ctxt i)  | 
| 
 
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
 | 
249  | 
ORELSE SOLVE_TIMEOUT (timeout div 20) "cvc3" (smt_solver_tac "cvc3" ctxt i)  | 
| 47812 | 250  | 
ORELSE SOLVE_TIMEOUT (timeout div 20) "best" (best_tac ctxt 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
 | 
251  | 
ORELSE SOLVE_TIMEOUT (timeout div 10) "force" (force_tac ctxt i)  | 
| 47832 | 252  | 
ORELSE SOLVE_TIMEOUT (timeout div 10) "meson" (Meson.meson_tac ctxt [] i)  | 
| 47812 | 253  | 
ORELSE SOLVE_TIMEOUT (timeout div 10) "fastforce" (fast_force_tac ctxt i)  | 
| 47771 | 254  | 
|
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
255  | 
fun problem_const_prefix thy = Context.theory_name thy ^ Long_Name.separator  | 
| 47785 | 256  | 
|
257  | 
(* Isabelle's standard automatic tactics ("auto", etc.) are more eager to
 | 
|
258  | 
unfold "definitions" of free variables than of constants (cf. PUZ107^5). *)  | 
|
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
259  | 
fun freeze_problem_consts thy =  | 
| 
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
260  | 
let val is_problem_const = String.isPrefix (problem_const_prefix thy) in  | 
| 
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
261  | 
map_aterms (fn t as Const (s, T) =>  | 
| 
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
262  | 
if is_problem_const s then Free (Long_Name.base_name s, T)  | 
| 
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
263  | 
else t  | 
| 
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
264  | 
| t => t)  | 
| 
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
265  | 
end  | 
| 47785 | 266  | 
|
| 
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
 | 
267  | 
fun make_conj (defs, nondefs) conjs =  | 
| 
 
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
 | 
268  | 
Logic.list_implies (rev defs @ rev nondefs,  | 
| 
 
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
 | 
269  | 
                      case conjs of conj :: _ => conj | [] => @{prop False})
 | 
| 47771 | 270  | 
|
271  | 
fun print_szs_from_success conjs success =  | 
|
| 47788 | 272  | 
  Output.urgent_message ("% SZS status " ^
 | 
273  | 
(if success then  | 
|
274  | 
if null conjs then "Unsatisfiable" else "Theorem"  | 
|
275  | 
else  | 
|
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
276  | 
"Unknown"))  | 
| 47765 | 277  | 
|
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
278  | 
fun sledgehammer_tptp_file thy timeout file_name =  | 
| 47771 | 279  | 
let  | 
| 47785 | 280  | 
val (conjs, assms, ctxt) =  | 
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
281  | 
read_tptp_file thy (freeze_problem_consts thy) file_name  | 
| 
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
 | 
282  | 
val conj = make_conj assms conjs  | 
| 47771 | 283  | 
in  | 
| 47832 | 284  | 
can_tac ctxt (sledgehammer_tac true ctxt timeout 1) conj  | 
| 47771 | 285  | 
|> print_szs_from_success conjs  | 
286  | 
end  | 
|
| 
47766
 
9f7cdd5fff7c
more work on TPTP Isabelle and Sledgehammer tactics
 
blanchet 
parents: 
47765 
diff
changeset
 | 
287  | 
|
| 
48083
 
a4148c95134d
renamed TPTP commands to agree with Sutcliffe's terminology
 
blanchet 
parents: 
48082 
diff
changeset
 | 
288  | 
fun generic_isabelle_tptp_file demo thy timeout file_name =  | 
| 47771 | 289  | 
let  | 
| 47785 | 290  | 
val (conjs, assms, ctxt) =  | 
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47793 
diff
changeset
 | 
291  | 
read_tptp_file thy (freeze_problem_consts thy) file_name  | 
| 
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
 | 
292  | 
val conj = make_conj assms conjs  | 
| 48143 | 293  | 
val (last_hope_atp, last_hope_completeness) =  | 
| 48082 | 294  | 
if demo then (ATP_Systems.satallaxN, 0) else (ATP_Systems.vampireN, 2)  | 
| 47771 | 295  | 
in  | 
| 47811 | 296  | 
(can_tac ctxt (auto_etc_tac ctxt (timeout div 2) 1) conj orelse  | 
| 47832 | 297  | 
can_tac ctxt (sledgehammer_tac demo ctxt (timeout div 2) 1) conj orelse  | 
| 48082 | 298  | 
can_tac ctxt (SOLVE_TIMEOUT timeout (last_hope_atp ^ "(*)")  | 
| 48143 | 299  | 
(atp_tac ctxt last_hope_completeness [] timeout last_hope_atp 1)) conj)  | 
| 47771 | 300  | 
|> print_szs_from_success conjs  | 
301  | 
end  | 
|
| 46324 | 302  | 
|
| 
48083
 
a4148c95134d
renamed TPTP commands to agree with Sutcliffe's terminology
 
blanchet 
parents: 
48082 
diff
changeset
 | 
303  | 
val isabelle_tptp_file = generic_isabelle_tptp_file false  | 
| 
 
a4148c95134d
renamed TPTP commands to agree with Sutcliffe's terminology
 
blanchet 
parents: 
48082 
diff
changeset
 | 
304  | 
val isabelle_hot_tptp_file = generic_isabelle_tptp_file true  | 
| 47832 | 305  | 
|
306  | 
||
| 46324 | 307  | 
(** Translator between TPTP(-like) file formats **)  | 
308  | 
||
| 46325 | 309  | 
fun translate_tptp_file format in_file_name out_file_name = ()  | 
| 46324 | 310  | 
|
311  | 
end;  |