| author | wenzelm | 
| Sat, 16 Oct 2021 20:21:13 +0200 | |
| changeset 74532 | 64d1b02327a4 | 
| parent 73684 | a63d76ba0a03 | 
| permissions | -rw-r--r-- | 
| 46324 | 1  | 
(* Title: HOL/TPTP/ATP_Problem_Import.thy  | 
2  | 
Author: Jasmin Blanchette, TU Muenchen  | 
|
3  | 
*)  | 
|
4  | 
||
| 63167 | 5  | 
section \<open>ATP Problem Importer\<close>  | 
| 47812 | 6  | 
|
| 46324 | 7  | 
theory ATP_Problem_Import  | 
| 
72001
 
3e08311ada8e
removed 'freeze_problem_consts' hack in TPTP tools, which wasn't compatible with post-2016 reforms to local theories
 
blanchet 
parents: 
69605 
diff
changeset
 | 
8  | 
imports Complex_Main TPTP_Interpret "HOL-Library.Refute"  | 
| 46324 | 9  | 
begin  | 
10  | 
||
| 63167 | 11  | 
ML \<open>Proofterm.proofs := 0\<close>  | 
| 47765 | 12  | 
|
| 47670 | 13  | 
declare [[show_consts]] (* for Refute *)  | 
| 47765 | 14  | 
declare [[smt_oracle]]  | 
| 47670 | 15  | 
|
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47790 
diff
changeset
 | 
16  | 
declare [[unify_search_bound = 60]]  | 
| 
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47790 
diff
changeset
 | 
17  | 
declare [[unify_trace_bound = 60]]  | 
| 
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47790 
diff
changeset
 | 
18  | 
|
| 69605 | 19  | 
ML_file \<open>atp_problem_import.ML\<close>  | 
| 
47794
 
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
 
blanchet 
parents: 
47790 
diff
changeset
 | 
20  | 
|
| 47785 | 21  | 
(*  | 
| 
57812
 
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
 
blanchet 
parents: 
52487 
diff
changeset
 | 
22  | 
ML {* ATP_Problem_Import.isabelle_tptp_file @{theory} 50 "$TPTP/Problems/PUZ/PUZ107^5.p" *}
 | 
| 64561 | 23  | 
ML {* ATP_Problem_Import.isabelle_hot_tptp_file @{theory} 20 "$TPTP/Problems/SYO/SYO304^5.p" *}
 | 
| 47785 | 24  | 
*)  | 
25  | 
||
| 46324 | 26  | 
end  |