|
1 theory Examples = Main: |
|
2 |
|
3 ML "reset eta_contract" |
|
4 ML "Pretty.setmargin 64" |
|
5 |
|
6 text{*membership, intersection *} |
|
7 text{*difference and empty set*} |
|
8 text{*complement, union and universal set*} |
|
9 |
|
10 lemma "(x \<in> A \<inter> B) = (x \<in> A \<and> x \<in> B)" |
|
11 apply (blast) |
|
12 done |
|
13 |
|
14 text{* |
|
15 @{thm[display] IntI[no_vars]} |
|
16 \rulename{IntI} |
|
17 |
|
18 @{thm[display] IntD1[no_vars]} |
|
19 \rulename{IntD1} |
|
20 |
|
21 @{thm[display] IntD2[no_vars]} |
|
22 \rulename{IntD2} |
|
23 *} |
|
24 |
|
25 lemma "(x \<in> -A) = (x \<notin> A)" |
|
26 apply (blast) |
|
27 done |
|
28 |
|
29 text{* |
|
30 @{thm[display] Compl_iff[no_vars]} |
|
31 \rulename{Compl_iff} |
|
32 *} |
|
33 |
|
34 lemma "- (A \<union> B) = -A \<inter> -B" |
|
35 apply (blast) |
|
36 done |
|
37 |
|
38 text{* |
|
39 @{thm[display] Compl_Un[no_vars]} |
|
40 \rulename{Compl_Un} |
|
41 *} |
|
42 |
|
43 lemma "A-A = {}" |
|
44 apply (blast) |
|
45 done |
|
46 |
|
47 text{* |
|
48 @{thm[display] Diff_disjoint[no_vars]} |
|
49 \rulename{Diff_disjoint} |
|
50 *} |
|
51 |
|
52 |
|
53 |
|
54 lemma "A \<union> -A = UNIV" |
|
55 apply (blast) |
|
56 done |
|
57 |
|
58 text{* |
|
59 @{thm[display] Compl_partition[no_vars]} |
|
60 \rulename{Compl_partition} |
|
61 *} |
|
62 |
|
63 text{*subset relation*} |
|
64 |
|
65 |
|
66 text{* |
|
67 @{thm[display] subsetI[no_vars]} |
|
68 \rulename{subsetI} |
|
69 |
|
70 @{thm[display] subsetD[no_vars]} |
|
71 \rulename{subsetD} |
|
72 *} |
|
73 |
|
74 lemma "((A \<union> B) \<subseteq> C) = (A \<subseteq> C \<and> B \<subseteq> C)" |
|
75 apply (blast) |
|
76 done |
|
77 |
|
78 text{* |
|
79 @{thm[display] Un_subset_iff[no_vars]} |
|
80 \rulename{Un_subset_iff} |
|
81 *} |
|
82 |
|
83 lemma "(A \<subseteq> -B) = (B \<subseteq> -A)" |
|
84 apply (blast) |
|
85 done |
|
86 |
|
87 lemma "(A <= -B) = (B <= -A)" |
|
88 oops |
|
89 |
|
90 text{*ASCII version: blast fails because of overloading because |
|
91 it doesn't have to be sets*} |
|
92 |
|
93 lemma "((A:: 'a set) <= -B) = (B <= -A)" |
|
94 apply (blast) |
|
95 done |
|
96 |
|
97 text{*A type constraint lets it work*} |
|
98 |
|
99 text{*An issue here: how do we discuss the distinction between ASCII and |
|
100 X-symbol notation? Here the latter disambiguates.*} |
|
101 |
|
102 |
|
103 text{* |
|
104 set extensionality |
|
105 |
|
106 @{thm[display] set_ext[no_vars]} |
|
107 \rulename{set_ext} |
|
108 |
|
109 @{thm[display] equalityI[no_vars]} |
|
110 \rulename{equalityI} |
|
111 |
|
112 @{thm[display] equalityE[no_vars]} |
|
113 \rulename{equalityE} |
|
114 *} |
|
115 |
|
116 |
|
117 text{*finite sets: insertion and membership relation*} |
|
118 text{*finite set notation*} |
|
119 |
|
120 lemma "insert x A = {x} \<union> A" |
|
121 apply (blast) |
|
122 done |
|
123 |
|
124 text{* |
|
125 @{thm[display] insert_is_Un[no_vars]} |
|
126 \rulename{insert_is_Un} |
|
127 *} |
|
128 |
|
129 lemma "{a,b} \<union> {c,d} = {a,b,c,d}" |
|
130 apply (blast) |
|
131 done |
|
132 |
|
133 lemma "{a,b} \<inter> {b,c} = {b}" |
|
134 apply (auto) |
|
135 oops |
|
136 text{*fails because it isn't valid*} |
|
137 |
|
138 lemma "{a,b} \<inter> {b,c} = (if a=c then {a,b} else {b})" |
|
139 apply (simp) |
|
140 apply (blast) |
|
141 done |
|
142 |
|
143 text{*or just force or auto. blast alone can't handle the if-then-else*} |
|
144 |
|
145 text{*next: some comprehension examples*} |
|
146 |
|
147 lemma "(a \<in> {z. P z}) = P a" |
|
148 apply (blast) |
|
149 done |
|
150 |
|
151 text{* |
|
152 @{thm[display] mem_Collect_eq[no_vars]} |
|
153 \rulename{mem_Collect_eq} |
|
154 *} |
|
155 |
|
156 lemma "{x. x \<in> A} = A" |
|
157 apply (blast) |
|
158 done |
|
159 |
|
160 text{* |
|
161 @{thm[display] Collect_mem_eq[no_vars]} |
|
162 \rulename{Collect_mem_eq} |
|
163 *} |
|
164 |
|
165 lemma "{x. P x \<or> x \<in> A} = {x. P x} \<union> A" |
|
166 apply (blast) |
|
167 done |
|
168 |
|
169 lemma "{x. P x \<longrightarrow> Q x} = -{x. P x} \<union> {x. Q x}" |
|
170 apply (blast) |
|
171 done |
|
172 |
|
173 constdefs |
|
174 prime :: "nat set" |
|
175 "prime == {p. 1<p & (ALL m. m dvd p --> m=1 | m=p)}" |
|
176 |
|
177 lemma "{p*q | p q. p\<in>prime \<and> q\<in>prime} = |
|
178 {z. \<exists>p q. z = p*q \<and> p\<in>prime \<and> q\<in>prime}" |
|
179 apply (blast) |
|
180 done |
|
181 |
|
182 text{*binders*} |
|
183 |
|
184 text{*bounded quantifiers*} |
|
185 |
|
186 lemma "(\<exists>x\<in>A. P x) = (\<exists>x. x\<in>A \<and> P x)" |
|
187 apply (blast) |
|
188 done |
|
189 |
|
190 text{* |
|
191 @{thm[display] bexI[no_vars]} |
|
192 \rulename{bexI} |
|
193 *} |
|
194 |
|
195 text{* |
|
196 @{thm[display] bexE[no_vars]} |
|
197 \rulename{bexE} |
|
198 *} |
|
199 |
|
200 lemma "(\<forall>x\<in>A. P x) = (\<forall>x. x\<in>A \<longrightarrow> P x)" |
|
201 apply (blast) |
|
202 done |
|
203 |
|
204 text{* |
|
205 @{thm[display] ballI[no_vars]} |
|
206 \rulename{ballI} |
|
207 *} |
|
208 |
|
209 text{* |
|
210 @{thm[display] bspec[no_vars]} |
|
211 \rulename{bspec} |
|
212 *} |
|
213 |
|
214 text{*indexed unions and variations*} |
|
215 |
|
216 lemma "(\<Union>x. B x) = (\<Union>x\<in>UNIV. B x)" |
|
217 apply (blast) |
|
218 done |
|
219 |
|
220 text{* |
|
221 @{thm[display] UN_iff[no_vars]} |
|
222 \rulename{UN_iff} |
|
223 *} |
|
224 |
|
225 text{* |
|
226 @{thm[display] Union_iff[no_vars]} |
|
227 \rulename{Union_iff} |
|
228 *} |
|
229 |
|
230 lemma "(\<Union>x\<in>A. B x) = {y. \<exists>x\<in>A. y \<in> B x}" |
|
231 apply (blast) |
|
232 done |
|
233 |
|
234 lemma "\<Union>S = (\<Union>x\<in>S. x)" |
|
235 apply (blast) |
|
236 done |
|
237 |
|
238 text{* |
|
239 @{thm[display] UN_I[no_vars]} |
|
240 \rulename{UN_I} |
|
241 *} |
|
242 |
|
243 text{* |
|
244 @{thm[display] UN_E[no_vars]} |
|
245 \rulename{UN_E} |
|
246 *} |
|
247 |
|
248 text{*indexed intersections*} |
|
249 |
|
250 lemma "(\<Inter>x. B x) = {y. \<forall>x. y \<in> B x}" |
|
251 apply (blast) |
|
252 done |
|
253 |
|
254 text{* |
|
255 @{thm[display] INT_iff[no_vars]} |
|
256 \rulename{INT_iff} |
|
257 *} |
|
258 |
|
259 text{* |
|
260 @{thm[display] Inter_iff[no_vars]} |
|
261 \rulename{Inter_iff} |
|
262 *} |
|
263 |
|
264 text{*mention also card, Pow, etc.*} |
|
265 |
|
266 |
|
267 text{* |
|
268 @{thm[display] card_Un_Int[no_vars]} |
|
269 \rulename{card_Un_Int} |
|
270 |
|
271 @{thm[display] card_Pow[no_vars]} |
|
272 \rulename{card_Pow} |
|
273 |
|
274 @{thm[display] n_subsets[no_vars]} |
|
275 \rulename{n_subsets} |
|
276 *} |
|
277 |
|
278 end |