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