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