| author | nipkow | 
| Mon, 13 Sep 2010 08:43:48 +0200 | |
| changeset 39301 | e1bd8a54c40f | 
| parent 38991 | 0e2798f30087 | 
| child 41144 | 509e51b7509a | 
| permissions | -rw-r--r-- | 
| 33027 | 1 | (* Title: HOL/Metis_Examples/Abstraction.thy | 
| 23449 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | ||
| 33027 | 4 | Testing the metis method. | 
| 23449 | 5 | *) | 
| 6 | ||
| 27368 | 7 | theory Abstraction | 
| 8 | imports Main FuncSet | |
| 23449 | 9 | begin | 
| 10 | ||
| 11 | (*For Christoph Benzmueller*) | |
| 12 | lemma "x<1 & ((op=) = (op=)) ==> ((op=) = (op=)) & (x<(2::nat))"; | |
| 13 | by (metis One_nat_def less_Suc0 not_less0 not_less_eq numeral_2_eq_2) | |
| 14 | ||
| 15 | (*this is a theorem, but we can't prove it unless ext is applied explicitly | |
| 16 | lemma "(op=) = (%x y. y=x)" | |
| 17 | *) | |
| 18 | ||
| 19 | consts | |
| 20 |   monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool"
 | |
| 21 | pset :: "'a set => 'a set" | |
| 22 |   order :: "'a set => ('a * 'a) set"
 | |
| 23 | ||
| 38991 | 24 | declare [[ sledgehammer_problem_prefix = "Abstraction__Collect_triv" ]] | 
| 23449 | 25 | lemma (*Collect_triv:*) "a \<in> {x. P x} ==> P a"
 | 
| 36566 | 26 | proof - | 
| 27 |   assume "a \<in> {x. P x}"
 | |
| 28 | hence "a \<in> P" by (metis Collect_def) | |
| 29 | hence "P a" by (metis mem_def) | |
| 30 | thus "P a" by metis | |
| 23449 | 31 | qed | 
| 32 | ||
| 33 | lemma Collect_triv: "a \<in> {x. P x} ==> P a"
 | |
| 23756 | 34 | by (metis mem_Collect_eq) | 
| 23449 | 35 | |
| 36 | ||
| 38991 | 37 | declare [[ sledgehammer_problem_prefix = "Abstraction__Collect_mp" ]] | 
| 23449 | 38 | lemma "a \<in> {x. P x --> Q x} ==> a \<in> {x. P x} ==> a \<in> {x. Q x}"
 | 
| 36566 | 39 | by (metis Collect_imp_eq ComplD UnE) | 
| 23449 | 40 | |
| 38991 | 41 | declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_triv" ]] | 
| 23449 | 42 | lemma "(a,b) \<in> Sigma A B ==> a \<in> A & b \<in> B a" | 
| 36566 | 43 | proof - | 
| 44 | assume A1: "(a, b) \<in> Sigma A B" | |
| 45 | hence F1: "b \<in> B a" by (metis mem_Sigma_iff) | |
| 46 | have F2: "a \<in> A" by (metis A1 mem_Sigma_iff) | |
| 47 | have "b \<in> B a" by (metis F1) | |
| 48 | thus "a \<in> A \<and> b \<in> B a" by (metis F2) | |
| 23449 | 49 | qed | 
| 50 | ||
| 51 | lemma Sigma_triv: "(a,b) \<in> Sigma A B ==> a \<in> A & b \<in> B a" | |
| 52 | by (metis SigmaD1 SigmaD2) | |
| 53 | ||
| 38991 | 54 | declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect" ]] | 
| 36566 | 55 | lemma "(a, b) \<in> (SIGMA x:A. {y. x = f y}) \<Longrightarrow> a \<in> A \<and> a = f b"
 | 
| 56 | (* Metis says this is satisfiable! | |
| 29676 | 57 | by (metis CollectD SigmaD1 SigmaD2) | 
| 58 | *) | |
| 23449 | 59 | by (meson CollectD SigmaD1 SigmaD2) | 
| 60 | ||
| 61 | ||
| 36566 | 62 | lemma "(a, b) \<in> (SIGMA x:A. {y. x = f y}) \<Longrightarrow> a \<in> A \<and> a = f b"
 | 
| 63 | by (metis mem_Sigma_iff singleton_conv2 vimage_Collect_eq vimage_singleton_eq) | |
| 24827 | 64 | |
| 36566 | 65 | lemma "(a, b) \<in> (SIGMA x:A. {y. x = f y}) \<Longrightarrow> a \<in> A \<and> a = f b"
 | 
| 66 | proof - | |
| 67 |   assume A1: "(a, b) \<in> (SIGMA x:A. {y. x = f y})"
 | |
| 68 |   have F1: "\<forall>u. {u} = op = u" by (metis singleton_conv2 Collect_def)
 | |
| 36571 | 69 | have F2: "\<forall>y w v. v \<in> w -` op = y \<longrightarrow> w v = y" | 
| 70 | by (metis F1 vimage_singleton_eq) | |
| 71 | have F3: "\<forall>x w. (\<lambda>R. w (x R)) = x -` w" | |
| 72 | by (metis vimage_Collect_eq Collect_def) | |
| 73 | show "a \<in> A \<and> a = f b" by (metis A1 F2 F3 mem_Sigma_iff Collect_def) | |
| 24827 | 74 | qed | 
| 75 | ||
| 36566 | 76 | (* Alternative structured proof *) | 
| 77 | lemma "(a, b) \<in> (SIGMA x:A. {y. x = f y}) \<Longrightarrow> a \<in> A \<and> a = f b"
 | |
| 78 | proof - | |
| 79 |   assume A1: "(a, b) \<in> (SIGMA x:A. {y. x = f y})"
 | |
| 80 | hence F1: "a \<in> A" by (metis mem_Sigma_iff) | |
| 81 |   have "b \<in> {R. a = f R}" by (metis A1 mem_Sigma_iff)
 | |
| 82 | hence F2: "b \<in> (\<lambda>R. a = f R)" by (metis Collect_def) | |
| 83 | hence "a = f b" by (unfold mem_def) | |
| 84 | thus "a \<in> A \<and> a = f b" by (metis F1) | |
| 24827 | 85 | qed | 
| 23449 | 86 | |
| 87 | ||
| 38991 | 88 | declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_in_pp" ]] | 
| 23449 | 89 | lemma "(cl,f) \<in> CLF ==> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) ==> f \<in> pset cl"
 | 
| 24827 | 90 | by (metis Collect_mem_eq SigmaD2) | 
| 23449 | 91 | |
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24632diff
changeset | 92 | lemma "(cl,f) \<in> CLF ==> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) ==> f \<in> pset cl"
 | 
| 36566 | 93 | proof - | 
| 94 | assume A1: "(cl, f) \<in> CLF" | |
| 95 |   assume A2: "CLF = (SIGMA cl:CL. {f. f \<in> pset cl})"
 | |
| 96 | have F1: "\<forall>v. (\<lambda>R. R \<in> v) = v" by (metis Collect_mem_eq Collect_def) | |
| 97 |   have "\<forall>v u. (u, v) \<in> CLF \<longrightarrow> v \<in> {R. R \<in> pset u}" by (metis A2 mem_Sigma_iff)
 | |
| 98 | hence "\<forall>v u. (u, v) \<in> CLF \<longrightarrow> v \<in> pset u" by (metis F1 Collect_def) | |
| 99 | hence "f \<in> pset cl" by (metis A1) | |
| 100 | thus "f \<in> pset cl" by metis | |
| 24827 | 101 | qed | 
| 23449 | 102 | |
| 38991 | 103 | declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect_Pi" ]] | 
| 23449 | 104 | lemma | 
| 105 |     "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==> 
 | |
| 106 | f \<in> pset cl \<rightarrow> pset cl" | |
| 36566 | 107 | proof - | 
| 108 |   assume A1: "(cl, f) \<in> (SIGMA cl:CL. {f. f \<in> pset cl \<rightarrow> pset cl})"
 | |
| 109 | have F1: "\<forall>v. (\<lambda>R. R \<in> v) = v" by (metis Collect_mem_eq Collect_def) | |
| 110 |   have "f \<in> {R. R \<in> pset cl \<rightarrow> pset cl}" using A1 by simp
 | |
| 111 | hence "f \<in> pset cl \<rightarrow> pset cl" by (metis F1 Collect_def) | |
| 112 | thus "f \<in> pset cl \<rightarrow> pset cl" by metis | |
| 24827 | 113 | qed | 
| 23449 | 114 | |
| 38991 | 115 | declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect_Int" ]] | 
| 23449 | 116 | lemma | 
| 117 |     "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
 | |
| 118 | f \<in> pset cl \<inter> cl" | |
| 36566 | 119 | proof - | 
| 120 |   assume A1: "(cl, f) \<in> (SIGMA cl:CL. {f. f \<in> pset cl \<inter> cl})"
 | |
| 121 | have F1: "\<forall>v. (\<lambda>R. R \<in> v) = v" by (metis Collect_mem_eq Collect_def) | |
| 122 |   have "f \<in> {R. R \<in> pset cl \<inter> cl}" using A1 by simp
 | |
| 123 | hence "f \<in> Id_on cl `` pset cl" by (metis F1 Int_commute Image_Id_on Collect_def) | |
| 124 | hence "f \<in> Id_on cl `` pset cl" by metis | |
| 125 | hence "f \<in> cl \<inter> pset cl" by (metis Image_Id_on) | |
| 126 | thus "f \<in> pset cl \<inter> cl" by (metis Int_commute) | |
| 24827 | 127 | qed | 
| 128 | ||
| 23449 | 129 | |
| 38991 | 130 | declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect_Pi_mono" ]] | 
| 23449 | 131 | lemma | 
| 132 |     "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) ==>
 | |
| 133 | (f \<in> pset cl \<rightarrow> pset cl) & (monotone f (pset cl) (order cl))" | |
| 134 | by auto | |
| 135 | ||
| 38991 | 136 | declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_subset_Collect_Int" ]] | 
| 23449 | 137 | lemma "(cl,f) \<in> CLF ==> | 
| 138 |    CLF \<subseteq> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
 | |
| 139 | f \<in> pset cl \<inter> cl" | |
| 24827 | 140 | by auto | 
| 27368 | 141 | |
| 23449 | 142 | |
| 38991 | 143 | declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_Collect_Int" ]] | 
| 23449 | 144 | lemma "(cl,f) \<in> CLF ==> | 
| 145 |    CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
 | |
| 146 | f \<in> pset cl \<inter> cl" | |
| 24827 | 147 | by auto | 
| 36566 | 148 | |
| 23449 | 149 | |
| 38991 | 150 | declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_subset_Collect_Pi" ]] | 
| 23449 | 151 | lemma | 
| 152 | "(cl,f) \<in> CLF ==> | |
| 153 |     CLF \<subseteq> (SIGMA cl': CL. {f. f \<in> pset cl' \<rightarrow> pset cl'}) ==> 
 | |
| 154 | f \<in> pset cl \<rightarrow> pset cl" | |
| 31754 | 155 | by fast | 
| 36566 | 156 | |
| 23449 | 157 | |
| 38991 | 158 | declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_Collect_Pi" ]] | 
| 23449 | 159 | lemma | 
| 160 | "(cl,f) \<in> CLF ==> | |
| 161 |    CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==> 
 | |
| 162 | f \<in> pset cl \<rightarrow> pset cl" | |
| 24827 | 163 | by auto | 
| 36566 | 164 | |
| 23449 | 165 | |
| 38991 | 166 | declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_Collect_Pi_mono" ]] | 
| 23449 | 167 | lemma | 
| 168 | "(cl,f) \<in> CLF ==> | |
| 169 |    CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) ==>
 | |
| 170 | (f \<in> pset cl \<rightarrow> pset cl) & (monotone f (pset cl) (order cl))" | |
| 171 | by auto | |
| 172 | ||
| 38991 | 173 | declare [[ sledgehammer_problem_prefix = "Abstraction__map_eq_zipA" ]] | 
| 23449 | 174 | lemma "map (%x. (f x, g x)) xs = zip (map f xs) (map g xs)" | 
| 175 | apply (induct xs) | |
| 36566 | 176 | apply (metis map_is_Nil_conv zip.simps(1)) | 
| 177 | by auto | |
| 23449 | 178 | |
| 38991 | 179 | declare [[ sledgehammer_problem_prefix = "Abstraction__map_eq_zipB" ]] | 
| 23449 | 180 | lemma "map (%w. (w -> w, w \<times> w)) xs = | 
| 181 | zip (map (%w. w -> w) xs) (map (%w. w \<times> w) xs)" | |
| 182 | apply (induct xs) | |
| 36566 | 183 | apply (metis Nil_is_map_conv zip_Nil) | 
| 184 | by auto | |
| 23449 | 185 | |
| 38991 | 186 | declare [[ sledgehammer_problem_prefix = "Abstraction__image_evenA" ]] | 
| 36566 | 187 | lemma "(%x. Suc(f x)) ` {x. even x} <= A ==> (\<forall>x. even x --> Suc(f x) \<in> A)"
 | 
| 188 | by (metis Collect_def image_subset_iff mem_def) | |
| 23449 | 189 | |
| 38991 | 190 | declare [[ sledgehammer_problem_prefix = "Abstraction__image_evenB" ]] | 
| 23449 | 191 | lemma "(%x. f (f x)) ` ((%x. Suc(f x)) ` {x. even x}) <= A 
 | 
| 192 | ==> (\<forall>x. even x --> f (f (Suc(f x))) \<in> A)"; | |
| 36566 | 193 | by (metis Collect_def imageI image_image image_subset_iff mem_def) | 
| 23449 | 194 | |
| 38991 | 195 | declare [[ sledgehammer_problem_prefix = "Abstraction__image_curry" ]] | 
| 23449 | 196 | lemma "f \<in> (%u v. b \<times> u \<times> v) ` A ==> \<forall>u v. P (b \<times> u \<times> v) ==> P(f y)" | 
| 36566 | 197 | (*sledgehammer*) | 
| 23449 | 198 | by auto | 
| 199 | ||
| 38991 | 200 | declare [[ sledgehammer_problem_prefix = "Abstraction__image_TimesA" ]] | 
| 23449 | 201 | lemma image_TimesA: "(%(x,y). (f x, g y)) ` (A \<times> B) = (f`A) \<times> (g`B)" | 
| 36566 | 202 | (*sledgehammer*) | 
| 23449 | 203 | apply (rule equalityI) | 
| 204 | (***Even the two inclusions are far too difficult | |
| 38991 | 205 | using [[ sledgehammer_problem_prefix = "Abstraction__image_TimesA_simpler"]] | 
| 23449 | 206 | ***) | 
| 207 | apply (rule subsetI) | |
| 208 | apply (erule imageE) | |
| 209 | (*V manages from here with help: Abstraction__image_TimesA_simpler_1_b.p*) | |
| 210 | apply (erule ssubst) | |
| 211 | apply (erule SigmaE) | |
| 212 | (*V manages from here: Abstraction__image_TimesA_simpler_1_a.p*) | |
| 213 | apply (erule ssubst) | |
| 214 | apply (subst split_conv) | |
| 215 | apply (rule SigmaI) | |
| 216 | apply (erule imageI) + | |
| 217 | txt{*subgoal 2*}
 | |
| 218 | apply (clarify ); | |
| 219 | apply (simp add: ); | |
| 220 | apply (rule rev_image_eqI) | |
| 221 | apply (blast intro: elim:); | |
| 222 | apply (simp add: ); | |
| 223 | done | |
| 224 | ||
| 225 | (*Given the difficulty of the previous problem, these two are probably | |
| 226 | impossible*) | |
| 227 | ||
| 38991 | 228 | declare [[ sledgehammer_problem_prefix = "Abstraction__image_TimesB" ]] | 
| 23449 | 229 | lemma image_TimesB: | 
| 36566 | 230 | "(%(x,y,z). (f x, g y, h z)) ` (A \<times> B \<times> C) = (f`A) \<times> (g`B) \<times> (h`C)" | 
| 231 | (*sledgehammer*) | |
| 23449 | 232 | by force | 
| 233 | ||
| 38991 | 234 | declare [[ sledgehammer_problem_prefix = "Abstraction__image_TimesC" ]] | 
| 23449 | 235 | lemma image_TimesC: | 
| 236 | "(%(x,y). (x \<rightarrow> x, y \<times> y)) ` (A \<times> B) = | |
| 237 | ((%x. x \<rightarrow> x) ` A) \<times> ((%y. y \<times> y) ` B)" | |
| 36566 | 238 | (*sledgehammer*) | 
| 23449 | 239 | by auto | 
| 240 | ||
| 241 | end |