| author | haftmann | 
| Wed, 11 Mar 2009 08:45:47 +0100 | |
| changeset 30430 | 42ea5d85edcc | 
| parent 29654 | 24e73987bfe2 | 
| child 31037 | ac8669134e7a | 
| permissions | -rw-r--r-- | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 1 | (* Title: HOL/ATP_Linkup.thy | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 2 | Author: Lawrence C Paulson | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 3 | Author: Jia Meng, NICTA | 
| 28483 | 4 | Author: Fabian Immler, TUM | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 5 | *) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 6 | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28476diff
changeset | 7 | header {* The Isabelle-ATP Linkup *}
 | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 8 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 9 | theory ATP_Linkup | 
| 29654 
24e73987bfe2
Plain, Main form meeting points in import hierarchy
 haftmann parents: 
29611diff
changeset | 10 | imports Divides Record Hilbert_Choice Plain | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 11 | uses | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 12 | "Tools/polyhash.ML" | 
| 21977 | 13 | "Tools/res_clause.ML" | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28476diff
changeset | 14 |   ("Tools/res_axioms.ML")
 | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 15 |   ("Tools/res_hol_clause.ML")
 | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21977diff
changeset | 16 |   ("Tools/res_reconstruct.ML")
 | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 17 |   ("Tools/res_atp.ML")
 | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28476diff
changeset | 18 |   ("Tools/atp_manager.ML")
 | 
| 28592 | 19 |   ("Tools/atp_wrapper.ML")
 | 
| 23444 | 20 | "~~/src/Tools/Metis/metis.ML" | 
| 21 |   ("Tools/metis_tools.ML")
 | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 22 | begin | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 23 | |
| 24819 | 24 | definition COMBI :: "'a => 'a" | 
| 25 | where "COMBI P == P" | |
| 26 | ||
| 27 | definition COMBK :: "'a => 'b => 'a" | |
| 28 | where "COMBK P Q == P" | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 29 | |
| 24819 | 30 | definition COMBB :: "('b => 'c) => ('a => 'b) => 'a => 'c"
 | 
| 31 | where "COMBB P Q R == P (Q R)" | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 32 | |
| 24819 | 33 | definition COMBC :: "('a => 'b => 'c) => 'b => 'a => 'c"
 | 
| 34 | where "COMBC P Q R == P R Q" | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 35 | |
| 24819 | 36 | definition COMBS :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c"
 | 
| 37 | where "COMBS P Q R == P R (Q R)" | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 38 | |
| 24819 | 39 | definition fequal :: "'a => 'a => bool" | 
| 40 | where "fequal X Y == (X=Y)" | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 41 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 42 | lemma fequal_imp_equal: "fequal X Y ==> X=Y" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 43 | by (simp add: fequal_def) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 44 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 45 | lemma equal_imp_fequal: "X=Y ==> fequal X Y" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 46 | by (simp add: fequal_def) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 47 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 48 | text{*These two represent the equivalence between Boolean equality and iff.
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 49 | They can't be converted to clauses automatically, as the iff would be | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 50 | expanded...*} | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 51 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 52 | lemma iff_positive: "P | Q | P=Q" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 53 | by blast | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 54 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 55 | lemma iff_negative: "~P | ~Q | P=Q" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 56 | by blast | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 57 | |
| 24827 | 58 | text{*Theorems for translation to combinators*}
 | 
| 59 | ||
| 60 | lemma abs_S: "(%x. (f x) (g x)) == COMBS f g" | |
| 61 | apply (rule eq_reflection) | |
| 62 | apply (rule ext) | |
| 63 | apply (simp add: COMBS_def) | |
| 64 | done | |
| 65 | ||
| 66 | lemma abs_I: "(%x. x) == COMBI" | |
| 67 | apply (rule eq_reflection) | |
| 68 | apply (rule ext) | |
| 69 | apply (simp add: COMBI_def) | |
| 70 | done | |
| 71 | ||
| 72 | lemma abs_K: "(%x. y) == COMBK y" | |
| 73 | apply (rule eq_reflection) | |
| 74 | apply (rule ext) | |
| 75 | apply (simp add: COMBK_def) | |
| 76 | done | |
| 77 | ||
| 78 | lemma abs_B: "(%x. a (g x)) == COMBB a g" | |
| 79 | apply (rule eq_reflection) | |
| 80 | apply (rule ext) | |
| 81 | apply (simp add: COMBB_def) | |
| 82 | done | |
| 83 | ||
| 84 | lemma abs_C: "(%x. (f x) b) == COMBC f b" | |
| 85 | apply (rule eq_reflection) | |
| 86 | apply (rule ext) | |
| 87 | apply (simp add: COMBC_def) | |
| 88 | done | |
| 89 | ||
| 27368 | 90 | |
| 28585 | 91 | subsection {* Setup of external ATPs *}
 | 
| 27368 | 92 | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28476diff
changeset | 93 | use "Tools/res_axioms.ML" setup ResAxioms.setup | 
| 24827 | 94 | use "Tools/res_hol_clause.ML" | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28476diff
changeset | 95 | use "Tools/res_reconstruct.ML" setup ResReconstruct.setup | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 96 | use "Tools/res_atp.ML" | 
| 28573 | 97 | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28476diff
changeset | 98 | use "Tools/atp_manager.ML" | 
| 28592 | 99 | use "Tools/atp_wrapper.ML" | 
| 28483 | 100 | |
| 101 | text {* basic provers *}
 | |
| 28592 | 102 | setup {* AtpManager.add_prover "spass" AtpWrapper.spass *}
 | 
| 103 | setup {* AtpManager.add_prover "vampire" AtpWrapper.vampire *}
 | |
| 104 | setup {* AtpManager.add_prover "e" AtpWrapper.eprover *}
 | |
| 28483 | 105 | |
| 106 | text {* provers with stuctured output *}
 | |
| 28592 | 107 | setup {* AtpManager.add_prover "vampire_full" AtpWrapper.vampire_full *}
 | 
| 108 | setup {* AtpManager.add_prover "e_full" AtpWrapper.eprover_full *}
 | |
| 28483 | 109 | |
| 110 | text {* on some problems better results *}
 | |
| 28594 | 111 | setup {* AtpManager.add_prover "spass_no_tc" (AtpWrapper.spass_opts 40 false) *}
 | 
| 27182 | 112 | |
| 28573 | 113 | text {* remote provers via SystemOnTPTP *}
 | 
| 29590 | 114 | setup {* AtpManager.add_prover "remote_vampire"
 | 
| 29593 | 115 | (AtpWrapper.remote_prover "-s Vampire---9.0") *} | 
| 29587 
96599d8d8268
simplified usage of remote-script; added compatible remote-atps
 immler@in.tum.de parents: 
28594diff
changeset | 116 | setup {* AtpManager.add_prover "remote_spass"
 | 
| 29593 | 117 | (AtpWrapper.remote_prover "-s SPASS---3.01") *} | 
| 29587 
96599d8d8268
simplified usage of remote-script; added compatible remote-atps
 immler@in.tum.de parents: 
28594diff
changeset | 118 | setup {* AtpManager.add_prover "remote_e"
 | 
| 29593 | 119 | (AtpWrapper.remote_prover "-s EP---1.0") *} | 
| 29587 
96599d8d8268
simplified usage of remote-script; added compatible remote-atps
 immler@in.tum.de parents: 
28594diff
changeset | 120 | |
| 28573 | 121 | |
| 23444 | 122 | |
| 123 | subsection {* The Metis prover *}
 | |
| 124 | ||
| 125 | use "Tools/metis_tools.ML" | |
| 126 | setup MetisTools.setup | |
| 127 | ||
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 128 | end |