author | wenzelm |
Mon, 11 Apr 2016 15:26:58 +0200 | |
changeset 62954 | c5d0fdc260fa |
parent 58889 | 5b7a9633cfa8 |
child 63167 | 0909deb8059b |
permissions | -rw-r--r-- |
46324 | 1 |
(* Title: HOL/TPTP/ATP_Problem_Import.thy |
2 |
Author: Jasmin Blanchette, TU Muenchen |
|
3 |
*) |
|
4 |
||
58889 | 5 |
section {* ATP Problem Importer *} |
47812 | 6 |
|
46324 | 7 |
theory ATP_Problem_Import |
57812
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
52487
diff
changeset
|
8 |
imports Complex_Main TPTP_Interpret "~~/src/HOL/Library/Refute" |
46324 | 9 |
begin |
10 |
||
48891 | 11 |
ML_file "sledgehammer_tactics.ML" |
12 |
||
52487
48bc24467008
backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
wenzelm
parents:
52470
diff
changeset
|
13 |
ML {* Proofterm.proofs := 0 *} |
47765 | 14 |
|
47670 | 15 |
declare [[show_consts]] (* for Refute *) |
47765 | 16 |
declare [[smt_oracle]] |
47670 | 17 |
|
47794
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
blanchet
parents:
47790
diff
changeset
|
18 |
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
|
19 |
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
|
20 |
|
48891 | 21 |
ML_file "atp_problem_import.ML" |
47794
4ad62c5f9f88
thread theory cleanly and use "smt" method rather than Sledgehammer for Z3 (because of obscure debilitating bug)
blanchet
parents:
47790
diff
changeset
|
22 |
|
47785 | 23 |
(* |
57812
8dc9dc241973
make TPTP tools work on polymorphic (TFF1) problems as well
blanchet
parents:
52487
diff
changeset
|
24 |
ML {* ATP_Problem_Import.isabelle_tptp_file @{theory} 50 "$TPTP/Problems/PUZ/PUZ107^5.p" *} |
47785 | 25 |
*) |
26 |
||
46324 | 27 |
end |