| author | haftmann | 
| Wed, 09 Jan 2008 08:32:09 +0100 | |
| changeset 25870 | a6a21adf3b55 | 
| parent 25741 | 2d102ddaca8b | 
| child 26729 | 43a72d892594 | 
| 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 | ID: $Id$ | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 3 | Author: Lawrence C Paulson | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 4 | Author: Jia Meng, NICTA | 
| 
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 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 7 | header{* The Isabelle-ATP Linkup *}
 | 
| 
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 | 
| 25591 
0792e02973cc
swtiched ATP_Linkup and PreList in theory hierarchy
 haftmann parents: 
24943diff
changeset | 10 | imports PreList Hilbert_Choice | 
| 25728 
71e33d95ac55
moved Pure/General/random_word.ML to Tools/random_word.ML;
 wenzelm parents: 
25710diff
changeset | 11 | (*FIXME It must be a parent or a child of every other theory, to prevent theory-merge errors. FIXME*) | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 12 | uses | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 13 | "Tools/polyhash.ML" | 
| 21977 | 14 | "Tools/res_clause.ML" | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 15 |   ("Tools/res_hol_clause.ML")
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 16 |   ("Tools/res_axioms.ML")
 | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21977diff
changeset | 17 |   ("Tools/res_reconstruct.ML")
 | 
| 23519 | 18 |   ("Tools/watcher.ML")
 | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 19 |   ("Tools/res_atp.ML")
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 20 |   ("Tools/res_atp_provers.ML")
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 21 |   ("Tools/res_atp_methods.ML")
 | 
| 23444 | 22 | "~~/src/Tools/Metis/metis.ML" | 
| 23 |   ("Tools/metis_tools.ML")
 | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 24 | begin | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 25 | |
| 24819 | 26 | definition COMBI :: "'a => 'a" | 
| 27 | where "COMBI P == P" | |
| 28 | ||
| 29 | definition COMBK :: "'a => 'b => 'a" | |
| 30 | where "COMBK P Q == P" | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 31 | |
| 24819 | 32 | definition COMBB :: "('b => 'c) => ('a => 'b) => 'a => 'c"
 | 
| 33 | where "COMBB P Q R == P (Q R)" | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 34 | |
| 24819 | 35 | definition COMBC :: "('a => 'b => 'c) => 'b => 'a => 'c"
 | 
| 36 | where "COMBC P Q R == P R Q" | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 37 | |
| 24819 | 38 | definition COMBS :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c"
 | 
| 39 | where "COMBS P Q R == P R (Q R)" | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 40 | |
| 24819 | 41 | definition fequal :: "'a => 'a => bool" | 
| 42 | where "fequal X Y == (X=Y)" | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 43 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 44 | lemma fequal_imp_equal: "fequal X Y ==> X=Y" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 45 | by (simp add: fequal_def) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 46 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 47 | lemma equal_imp_fequal: "X=Y ==> fequal X Y" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 48 | by (simp add: fequal_def) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 49 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 50 | 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 | 51 | 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 | 52 | expanded...*} | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 53 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 54 | lemma iff_positive: "P | Q | P=Q" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 55 | by blast | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 56 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 57 | lemma iff_negative: "~P | ~Q | P=Q" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 58 | by blast | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 59 | |
| 24827 | 60 | text{*Theorems for translation to combinators*}
 | 
| 61 | ||
| 62 | lemma abs_S: "(%x. (f x) (g x)) == COMBS f g" | |
| 63 | apply (rule eq_reflection) | |
| 64 | apply (rule ext) | |
| 65 | apply (simp add: COMBS_def) | |
| 66 | done | |
| 67 | ||
| 68 | lemma abs_I: "(%x. x) == COMBI" | |
| 69 | apply (rule eq_reflection) | |
| 70 | apply (rule ext) | |
| 71 | apply (simp add: COMBI_def) | |
| 72 | done | |
| 73 | ||
| 74 | lemma abs_K: "(%x. y) == COMBK y" | |
| 75 | apply (rule eq_reflection) | |
| 76 | apply (rule ext) | |
| 77 | apply (simp add: COMBK_def) | |
| 78 | done | |
| 79 | ||
| 80 | lemma abs_B: "(%x. a (g x)) == COMBB a g" | |
| 81 | apply (rule eq_reflection) | |
| 82 | apply (rule ext) | |
| 83 | apply (simp add: COMBB_def) | |
| 84 | done | |
| 85 | ||
| 86 | lemma abs_C: "(%x. (f x) b) == COMBC f b" | |
| 87 | apply (rule eq_reflection) | |
| 88 | apply (rule ext) | |
| 89 | apply (simp add: COMBC_def) | |
| 90 | done | |
| 91 | ||
| 92 | ||
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21977diff
changeset | 93 | use "Tools/res_axioms.ML"      --{*requires the combinators declared above*}
 | 
| 24827 | 94 | use "Tools/res_hol_clause.ML" | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21977diff
changeset | 95 | use "Tools/res_reconstruct.ML" | 
| 23519 | 96 | use "Tools/watcher.ML" | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 97 | use "Tools/res_atp.ML" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 98 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 99 | setup ResAxioms.meson_method_setup | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 100 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 101 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 102 | subsection {* Setup for Vampire, E prover and SPASS *}
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 103 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 104 | use "Tools/res_atp_provers.ML" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 105 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 106 | oracle vampire_oracle ("string * int") = {* ResAtpProvers.vampire_o *}
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 107 | oracle eprover_oracle ("string * int") = {* ResAtpProvers.eprover_o *}
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 108 | oracle spass_oracle ("string * int") = {* ResAtpProvers.spass_o *}
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 109 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 110 | use "Tools/res_atp_methods.ML" | 
| 24827 | 111 | setup ResAtpMethods.setup      --{*Oracle ATP methods: still useful?*}
 | 
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25591diff
changeset | 112 | setup ResReconstruct.setup     --{*Config parameters*}
 | 
| 24827 | 113 | setup ResAxioms.setup          --{*Sledgehammer*}
 | 
| 23444 | 114 | |
| 115 | subsection {* The Metis prover *}
 | |
| 116 | ||
| 117 | use "Tools/metis_tools.ML" | |
| 118 | setup MetisTools.setup | |
| 119 | ||
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24318diff
changeset | 120 | setup {*
 | 
| 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24318diff
changeset | 121 | Theory.at_end ResAxioms.clause_cache_endtheory | 
| 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24318diff
changeset | 122 | *} | 
| 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24318diff
changeset | 123 | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 124 | end |