| author | wenzelm | 
| Thu, 20 Sep 2007 20:56:33 +0200 | |
| changeset 24665 | e5bea50b9b89 | 
| parent 24318 | 2477286fcc7e | 
| child 24742 | 73b8b42a36b6 | 
| 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 | 
| 24161 | 10 | imports Divides Record Hilbert_Choice | 
| 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" | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 14 |   ("Tools/res_hol_clause.ML")
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 15 |   ("Tools/res_axioms.ML")
 | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21977diff
changeset | 16 |   ("Tools/res_reconstruct.ML")
 | 
| 23519 | 17 |   ("Tools/watcher.ML")
 | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 18 |   ("Tools/res_atp.ML")
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 19 |   ("Tools/res_atp_provers.ML")
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 20 |   ("Tools/res_atp_methods.ML")
 | 
| 23444 | 21 | "~~/src/Tools/Metis/metis.ML" | 
| 22 |   ("Tools/metis_tools.ML")
 | |
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 23 | begin | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 24 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 25 | constdefs | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 26 | COMBI :: "'a => 'a" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 27 | "COMBI P == P" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 28 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 29 | COMBK :: "'a => 'b => 'a" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 30 | "COMBK P Q == P" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 31 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 32 |   COMBB :: "('b => 'c) => ('a => 'b) => 'a => 'c"
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 33 | "COMBB P Q R == P (Q R)" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 34 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 35 |   COMBC :: "('a => 'b => 'c) => 'b => 'a => 'c"
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 36 | "COMBC P Q R == P R Q" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 37 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 38 |   COMBS :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c"
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 39 | "COMBS P Q R == P R (Q R)" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 40 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 41 |   COMBB' :: "('a => 'c) => ('b => 'a) => ('d => 'b) => 'd => 'c"
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 42 | "COMBB' M P Q R == M (P (Q R))" | 
| 
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 |   COMBC' :: "('a => 'b => 'c) => ('d => 'a) => 'b => 'd => 'c"
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 45 | "COMBC' M P Q R == M (P R) Q" | 
| 
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 |   COMBS' :: "('a => 'b => 'c) => ('d => 'a) => ('d => 'b) => 'd => 'c"
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 48 | "COMBS' M P Q R == M (P R) (Q R)" | 
| 
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 | fequal :: "'a => 'a => bool" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 51 | "fequal X Y == (X=Y)" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 52 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 53 | lemma fequal_imp_equal: "fequal X Y ==> X=Y" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 54 | by (simp add: fequal_def) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 55 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 56 | lemma equal_imp_fequal: "X=Y ==> fequal X Y" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 57 | by (simp add: fequal_def) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 58 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 59 | lemma K_simp: "COMBK P == (%Q. P)" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 60 | apply (rule eq_reflection) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 61 | apply (rule ext) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 62 | apply (simp add: COMBK_def) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 63 | done | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 64 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 65 | lemma I_simp: "COMBI == (%P. P)" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 66 | apply (rule eq_reflection) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 67 | apply (rule ext) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 68 | apply (simp add: COMBI_def) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 69 | done | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 70 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 71 | lemma B_simp: "COMBB P Q == %R. P (Q R)" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 72 | apply (rule eq_reflection) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 73 | apply (rule ext) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 74 | apply (simp add: COMBB_def) | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 75 | done | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 76 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 77 | 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 | 78 | 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 | 79 | expanded...*} | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 80 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 81 | lemma iff_positive: "P | Q | P=Q" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 82 | by blast | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 83 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 84 | lemma iff_negative: "~P | ~Q | P=Q" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 85 | by blast | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 86 | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21977diff
changeset | 87 | use "Tools/res_axioms.ML"      --{*requires the combinators declared above*}
 | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21977diff
changeset | 88 | use "Tools/res_hol_clause.ML"  --{*requires the combinators*}
 | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21977diff
changeset | 89 | use "Tools/res_reconstruct.ML" | 
| 23519 | 90 | use "Tools/watcher.ML" | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 91 | use "Tools/res_atp.ML" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 92 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 93 | setup ResAxioms.meson_method_setup | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 94 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 95 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 96 | subsection {* Setup for Vampire, E prover and SPASS *}
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 97 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 98 | use "Tools/res_atp_provers.ML" | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 99 | |
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 100 | oracle vampire_oracle ("string * int") = {* ResAtpProvers.vampire_o *}
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 101 | oracle eprover_oracle ("string * int") = {* ResAtpProvers.eprover_o *}
 | 
| 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 102 | oracle spass_oracle ("string * int") = {* ResAtpProvers.spass_o *}
 | 
| 
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_methods.ML" | 
| 24318 | 105 | setup ResAtpMethods.setup | 
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 106 | |
| 23444 | 107 | |
| 108 | subsection {* The Metis prover *}
 | |
| 109 | ||
| 110 | use "Tools/metis_tools.ML" | |
| 111 | setup MetisTools.setup | |
| 112 | ||
| 21254 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 wenzelm parents: diff
changeset | 113 | end |