author | paulson |
Tue, 06 Mar 2012 15:15:49 +0000 | |
changeset 46820 | c656222c4dc1 |
parent 45620 | f2a587696afb |
child 46821 | ff6b0c1087f2 |
permissions | -rw-r--r-- |
2469 | 1 |
(* Title: ZF/upair.thy |
2 |
Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory |
|
3 |
Copyright 1993 University of Cambridge |
|
13259 | 4 |
|
5 |
Observe the order of dependence: |
|
6 |
Upair is defined in terms of Replace |
|
46820 | 7 |
\<union> is defined in terms of Upair and \<Union>(similarly for Int) |
13259 | 8 |
cons is defined in terms of Upair and Un |
9 |
Ordered pairs and descriptions are defined using cons ("set notation") |
|
2469 | 10 |
*) |
11 |
||
13357 | 12 |
header{*Unordered Pairs*} |
13 |
||
16417 | 14 |
theory upair imports ZF |
15 |
uses "Tools/typechk.ML" begin |
|
6153 | 16 |
|
9907 | 17 |
setup TypeCheck.setup |
6153 | 18 |
|
13780 | 19 |
lemma atomize_ball [symmetric, rulify]: |
46820 | 20 |
"(!!x. x:A ==> P(x)) == Trueprop (\<forall>x\<in>A. P(x))" |
13780 | 21 |
by (simp add: Ball_def atomize_all atomize_imp) |
13259 | 22 |
|
23 |
||
13357 | 24 |
subsection{*Unordered Pairs: constant @{term Upair}*} |
13259 | 25 |
|
46820 | 26 |
lemma Upair_iff [simp]: "c \<in> Upair(a,b) <-> (c=a | c=b)" |
13259 | 27 |
by (unfold Upair_def, blast) |
28 |
||
46820 | 29 |
lemma UpairI1: "a \<in> Upair(a,b)" |
13259 | 30 |
by simp |
31 |
||
46820 | 32 |
lemma UpairI2: "b \<in> Upair(a,b)" |
13259 | 33 |
by simp |
34 |
||
46820 | 35 |
lemma UpairE: "[| a \<in> Upair(b,c); a=b ==> P; a=c ==> P |] ==> P" |
13780 | 36 |
by (simp, blast) |
13259 | 37 |
|
13357 | 38 |
subsection{*Rules for Binary Union, Defined via @{term Upair}*} |
13259 | 39 |
|
46820 | 40 |
lemma Un_iff [simp]: "c \<in> A \<union> B <-> (c:A | c:B)" |
13259 | 41 |
apply (simp add: Un_def) |
42 |
apply (blast intro: UpairI1 UpairI2 elim: UpairE) |
|
43 |
done |
|
44 |
||
46820 | 45 |
lemma UnI1: "c \<in> A ==> c \<in> A \<union> B" |
13259 | 46 |
by simp |
47 |
||
46820 | 48 |
lemma UnI2: "c \<in> B ==> c \<in> A \<union> B" |
13259 | 49 |
by simp |
50 |
||
13356 | 51 |
declare UnI1 [elim?] UnI2 [elim?] |
52 |
||
46820 | 53 |
lemma UnE [elim!]: "[| c \<in> A \<union> B; c:A ==> P; c:B ==> P |] ==> P" |
13780 | 54 |
by (simp, blast) |
13259 | 55 |
|
56 |
(*Stronger version of the rule above*) |
|
46820 | 57 |
lemma UnE': "[| c \<in> A \<union> B; c:A ==> P; [| c:B; c\<notin>A |] ==> P |] ==> P" |
13780 | 58 |
by (simp, blast) |
13259 | 59 |
|
60 |
(*Classical introduction rule: no commitment to A vs B*) |
|
46820 | 61 |
lemma UnCI [intro!]: "(c \<notin> B ==> c \<in> A) ==> c \<in> A \<union> B" |
13780 | 62 |
by (simp, blast) |
13259 | 63 |
|
13357 | 64 |
subsection{*Rules for Binary Intersection, Defined via @{term Upair}*} |
13259 | 65 |
|
46820 | 66 |
lemma Int_iff [simp]: "c \<in> A \<inter> B <-> (c:A & c:B)" |
13259 | 67 |
apply (unfold Int_def) |
68 |
apply (blast intro: UpairI1 UpairI2 elim: UpairE) |
|
69 |
done |
|
70 |
||
46820 | 71 |
lemma IntI [intro!]: "[| c \<in> A; c \<in> B |] ==> c \<in> A \<inter> B" |
13259 | 72 |
by simp |
73 |
||
46820 | 74 |
lemma IntD1: "c \<in> A \<inter> B ==> c \<in> A" |
13259 | 75 |
by simp |
76 |
||
46820 | 77 |
lemma IntD2: "c \<in> A \<inter> B ==> c \<in> B" |
13259 | 78 |
by simp |
79 |
||
46820 | 80 |
lemma IntE [elim!]: "[| c \<in> A \<inter> B; [| c:A; c:B |] ==> P |] ==> P" |
13259 | 81 |
by simp |
82 |
||
83 |
||
13357 | 84 |
subsection{*Rules for Set Difference, Defined via @{term Upair}*} |
13259 | 85 |
|
46820 | 86 |
lemma Diff_iff [simp]: "c \<in> A-B <-> (c:A & c\<notin>B)" |
13259 | 87 |
by (unfold Diff_def, blast) |
88 |
||
46820 | 89 |
lemma DiffI [intro!]: "[| c \<in> A; c \<notin> B |] ==> c \<in> A - B" |
13259 | 90 |
by simp |
91 |
||
46820 | 92 |
lemma DiffD1: "c \<in> A - B ==> c \<in> A" |
13259 | 93 |
by simp |
94 |
||
46820 | 95 |
lemma DiffD2: "c \<in> A - B ==> c \<notin> B" |
13259 | 96 |
by simp |
97 |
||
46820 | 98 |
lemma DiffE [elim!]: "[| c \<in> A - B; [| c:A; c\<notin>B |] ==> P |] ==> P" |
13259 | 99 |
by simp |
100 |
||
101 |
||
13357 | 102 |
subsection{*Rules for @{term cons}*} |
13259 | 103 |
|
46820 | 104 |
lemma cons_iff [simp]: "a \<in> cons(b,A) <-> (a=b | a:A)" |
13259 | 105 |
apply (unfold cons_def) |
106 |
apply (blast intro: UpairI1 UpairI2 elim: UpairE) |
|
107 |
done |
|
108 |
||
109 |
(*risky as a typechecking rule, but solves otherwise unconstrained goals of |
|
46820 | 110 |
the form x \<in> ?A*) |
111 |
lemma consI1 [simp,TC]: "a \<in> cons(a,B)" |
|
13259 | 112 |
by simp |
113 |
||
114 |
||
46820 | 115 |
lemma consI2: "a \<in> B ==> a \<in> cons(b,B)" |
13259 | 116 |
by simp |
117 |
||
46820 | 118 |
lemma consE [elim!]: "[| a \<in> cons(b,A); a=b ==> P; a:A ==> P |] ==> P" |
13780 | 119 |
by (simp, blast) |
13259 | 120 |
|
121 |
(*Stronger version of the rule above*) |
|
122 |
lemma consE': |
|
46820 | 123 |
"[| a \<in> cons(b,A); a=b ==> P; [| a:A; a\<noteq>b |] ==> P |] ==> P" |
13780 | 124 |
by (simp, blast) |
13259 | 125 |
|
126 |
(*Classical introduction rule*) |
|
46820 | 127 |
lemma consCI [intro!]: "(a\<notin>B ==> a=b) ==> a: cons(b,B)" |
13780 | 128 |
by (simp, blast) |
13259 | 129 |
|
46820 | 130 |
lemma cons_not_0 [simp]: "cons(a,B) \<noteq> 0" |
13259 | 131 |
by (blast elim: equalityE) |
132 |
||
45602 | 133 |
lemmas cons_neq_0 = cons_not_0 [THEN notE] |
13259 | 134 |
|
135 |
declare cons_not_0 [THEN not_sym, simp] |
|
136 |
||
137 |
||
13357 | 138 |
subsection{*Singletons*} |
13259 | 139 |
|
46820 | 140 |
lemma singleton_iff: "a \<in> {b} <-> a=b" |
13259 | 141 |
by simp |
142 |
||
46820 | 143 |
lemma singletonI [intro!]: "a \<in> {a}" |
13259 | 144 |
by (rule consI1) |
145 |
||
45602 | 146 |
lemmas singletonE = singleton_iff [THEN iffD1, elim_format, elim!] |
13259 | 147 |
|
148 |
||
14883 | 149 |
subsection{*Descriptions*} |
13259 | 150 |
|
151 |
lemma the_equality [intro]: |
|
152 |
"[| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a" |
|
46820 | 153 |
apply (unfold the_def) |
13259 | 154 |
apply (fast dest: subst) |
155 |
done |
|
156 |
||
157 |
(* Only use this if you already know EX!x. P(x) *) |
|
158 |
lemma the_equality2: "[| EX! x. P(x); P(a) |] ==> (THE x. P(x)) = a" |
|
159 |
by blast |
|
160 |
||
161 |
lemma theI: "EX! x. P(x) ==> P(THE x. P(x))" |
|
162 |
apply (erule ex1E) |
|
163 |
apply (subst the_equality) |
|
164 |
apply (blast+) |
|
165 |
done |
|
166 |
||
46820 | 167 |
(*No congruence rule is necessary: if @{term"\<forall>y.P(y)<->Q(y)"} then |
168 |
@{term "THE x.P(x)"} rewrites to @{term "THE x.Q(x)"} *) |
|
13259 | 169 |
|
170 |
(*If it's "undefined", it's zero!*) |
|
171 |
lemma the_0: "~ (EX! x. P(x)) ==> (THE x. P(x))=0" |
|
172 |
apply (unfold the_def) |
|
173 |
apply (blast elim!: ReplaceE) |
|
174 |
done |
|
175 |
||
176 |
(*Easier to apply than theI: conclusion has only one occurrence of P*) |
|
177 |
lemma theI2: |
|
178 |
assumes p1: "~ Q(0) ==> EX! x. P(x)" |
|
179 |
and p2: "!!x. P(x) ==> Q(x)" |
|
180 |
shows "Q(THE x. P(x))" |
|
181 |
apply (rule classical) |
|
182 |
apply (rule p2) |
|
183 |
apply (rule theI) |
|
184 |
apply (rule classical) |
|
185 |
apply (rule p1) |
|
186 |
apply (erule the_0 [THEN subst], assumption) |
|
187 |
done |
|
188 |
||
13357 | 189 |
lemma the_eq_trivial [simp]: "(THE x. x = a) = a" |
190 |
by blast |
|
191 |
||
13544 | 192 |
lemma the_eq_trivial2 [simp]: "(THE x. a = x) = a" |
193 |
by blast |
|
194 |
||
13780 | 195 |
|
13357 | 196 |
subsection{*Conditional Terms: @{text "if-then-else"}*} |
13259 | 197 |
|
198 |
lemma if_true [simp]: "(if True then a else b) = a" |
|
199 |
by (unfold if_def, blast) |
|
200 |
||
201 |
lemma if_false [simp]: "(if False then a else b) = b" |
|
202 |
by (unfold if_def, blast) |
|
203 |
||
204 |
(*Never use with case splitting, or if P is known to be true or false*) |
|
205 |
lemma if_cong: |
|
46820 | 206 |
"[| P<->Q; Q ==> a=c; ~Q ==> b=d |] |
13259 | 207 |
==> (if P then a else b) = (if Q then c else d)" |
208 |
by (simp add: if_def cong add: conj_cong) |
|
209 |
||
210 |
(*Prevents simplification of x and y: faster and allows the execution |
|
211 |
of functional programs. NOW THE DEFAULT.*) |
|
212 |
lemma if_weak_cong: "P<->Q ==> (if P then x else y) = (if Q then x else y)" |
|
213 |
by simp |
|
214 |
||
215 |
(*Not needed for rewriting, since P would rewrite to True anyway*) |
|
216 |
lemma if_P: "P ==> (if P then a else b) = a" |
|
217 |
by (unfold if_def, blast) |
|
218 |
||
219 |
(*Not needed for rewriting, since P would rewrite to False anyway*) |
|
220 |
lemma if_not_P: "~P ==> (if P then a else b) = b" |
|
221 |
by (unfold if_def, blast) |
|
222 |
||
13780 | 223 |
lemma split_if [split]: |
46820 | 224 |
"P(if Q then x else y) <-> ((Q \<longrightarrow> P(x)) & (~Q \<longrightarrow> P(y)))" |
14153 | 225 |
by (case_tac Q, simp_all) |
13259 | 226 |
|
45620
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45602
diff
changeset
|
227 |
(** Rewrite rules for boolean case-splitting: faster than split_if [split] |
13259 | 228 |
**) |
229 |
||
45602 | 230 |
lemmas split_if_eq1 = split_if [of "%x. x = b"] for b |
231 |
lemmas split_if_eq2 = split_if [of "%x. a = x"] for x |
|
13259 | 232 |
|
46820 | 233 |
lemmas split_if_mem1 = split_if [of "%x. x \<in> b"] for b |
234 |
lemmas split_if_mem2 = split_if [of "%x. a \<in> x"] for x |
|
13259 | 235 |
|
236 |
lemmas split_ifs = split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2 |
|
237 |
||
238 |
(*Logically equivalent to split_if_mem2*) |
|
239 |
lemma if_iff: "a: (if P then x else y) <-> P & a:x | ~P & a:y" |
|
13780 | 240 |
by simp |
13259 | 241 |
|
242 |
lemma if_type [TC]: |
|
243 |
"[| P ==> a: A; ~P ==> b: A |] ==> (if P then a else b): A" |
|
13780 | 244 |
by simp |
245 |
||
246 |
(** Splitting IFs in the assumptions **) |
|
247 |
||
248 |
lemma split_if_asm: "P(if Q then x else y) <-> (~((Q & ~P(x)) | (~Q & ~P(y))))" |
|
249 |
by simp |
|
250 |
||
251 |
lemmas if_splits = split_if split_if_asm |
|
13259 | 252 |
|
253 |
||
13357 | 254 |
subsection{*Consequences of Foundation*} |
13259 | 255 |
|
256 |
(*was called mem_anti_sym*) |
|
257 |
lemma mem_asym: "[| a:b; ~P ==> b:a |] ==> P" |
|
258 |
apply (rule classical) |
|
259 |
apply (rule_tac A1 = "{a,b}" in foundation [THEN disjE]) |
|
260 |
apply (blast elim!: equalityE)+ |
|
261 |
done |
|
262 |
||
263 |
(*was called mem_anti_refl*) |
|
264 |
lemma mem_irrefl: "a:a ==> P" |
|
265 |
by (blast intro: mem_asym) |
|
266 |
||
267 |
(*mem_irrefl should NOT be added to default databases: |
|
268 |
it would be tried on most goals, making proofs slower!*) |
|
269 |
||
46820 | 270 |
lemma mem_not_refl: "a \<notin> a" |
13259 | 271 |
apply (rule notI) |
272 |
apply (erule mem_irrefl) |
|
273 |
done |
|
274 |
||
275 |
(*Good for proving inequalities by rewriting*) |
|
46820 | 276 |
lemma mem_imp_not_eq: "a:A ==> a \<noteq> A" |
13259 | 277 |
by (blast elim!: mem_irrefl) |
278 |
||
46820 | 279 |
lemma eq_imp_not_mem: "a=A ==> a \<notin> A" |
13357 | 280 |
by (blast intro: elim: mem_irrefl) |
281 |
||
282 |
subsection{*Rules for Successor*} |
|
13259 | 283 |
|
46820 | 284 |
lemma succ_iff: "i \<in> succ(j) <-> i=j | i:j" |
13259 | 285 |
by (unfold succ_def, blast) |
286 |
||
46820 | 287 |
lemma succI1 [simp]: "i \<in> succ(i)" |
13259 | 288 |
by (simp add: succ_iff) |
289 |
||
46820 | 290 |
lemma succI2: "i \<in> j ==> i \<in> succ(j)" |
13259 | 291 |
by (simp add: succ_iff) |
292 |
||
46820 | 293 |
lemma succE [elim!]: |
294 |
"[| i \<in> succ(j); i=j ==> P; i:j ==> P |] ==> P" |
|
295 |
apply (simp add: succ_iff, blast) |
|
13259 | 296 |
done |
297 |
||
298 |
(*Classical introduction rule*) |
|
46820 | 299 |
lemma succCI [intro!]: "(i\<notin>j ==> i=j) ==> i: succ(j)" |
13259 | 300 |
by (simp add: succ_iff, blast) |
301 |
||
46820 | 302 |
lemma succ_not_0 [simp]: "succ(n) \<noteq> 0" |
13259 | 303 |
by (blast elim!: equalityE) |
304 |
||
45602 | 305 |
lemmas succ_neq_0 = succ_not_0 [THEN notE, elim!] |
13259 | 306 |
|
307 |
declare succ_not_0 [THEN not_sym, simp] |
|
308 |
declare sym [THEN succ_neq_0, elim!] |
|
309 |
||
46820 | 310 |
(* @{term"succ(c) \<subseteq> B ==> c \<in> B"} *) |
13259 | 311 |
lemmas succ_subsetD = succI1 [THEN [2] subsetD] |
312 |
||
46820 | 313 |
(* @{term"succ(b) \<noteq> b"} *) |
45602 | 314 |
lemmas succ_neq_self = succI1 [THEN mem_imp_not_eq, THEN not_sym] |
13259 | 315 |
|
316 |
lemma succ_inject_iff [simp]: "succ(m) = succ(n) <-> m=n" |
|
317 |
by (blast elim: mem_asym elim!: equalityE) |
|
318 |
||
45602 | 319 |
lemmas succ_inject = succ_inject_iff [THEN iffD1, dest!] |
13259 | 320 |
|
13780 | 321 |
|
322 |
subsection{*Miniscoping of the Bounded Universal Quantifier*} |
|
323 |
||
324 |
lemma ball_simps1: |
|
46820 | 325 |
"(\<forall>x\<in>A. P(x) & Q) <-> (\<forall>x\<in>A. P(x)) & (A=0 | Q)" |
326 |
"(\<forall>x\<in>A. P(x) | Q) <-> ((\<forall>x\<in>A. P(x)) | Q)" |
|
327 |
"(\<forall>x\<in>A. P(x) \<longrightarrow> Q) <-> ((\<exists>x\<in>A. P(x)) \<longrightarrow> Q)" |
|
328 |
"(~(\<forall>x\<in>A. P(x))) <-> (\<exists>x\<in>A. ~P(x))" |
|
329 |
"(\<forall>x\<in>0.P(x)) <-> True" |
|
330 |
"(\<forall>x\<in>succ(i).P(x)) <-> P(i) & (\<forall>x\<in>i. P(x))" |
|
331 |
"(\<forall>x\<in>cons(a,B).P(x)) <-> P(a) & (\<forall>x\<in>B. P(x))" |
|
332 |
"(\<forall>x\<in>RepFun(A,f). P(x)) <-> (\<forall>y\<in>A. P(f(y)))" |
|
333 |
"(\<forall>x\<in>\<Union>(A).P(x)) <-> (\<forall>y\<in>A. \<forall>x\<in>y. P(x))" |
|
13780 | 334 |
by blast+ |
335 |
||
336 |
lemma ball_simps2: |
|
46820 | 337 |
"(\<forall>x\<in>A. P & Q(x)) <-> (A=0 | P) & (\<forall>x\<in>A. Q(x))" |
338 |
"(\<forall>x\<in>A. P | Q(x)) <-> (P | (\<forall>x\<in>A. Q(x)))" |
|
339 |
"(\<forall>x\<in>A. P \<longrightarrow> Q(x)) <-> (P \<longrightarrow> (\<forall>x\<in>A. Q(x)))" |
|
13780 | 340 |
by blast+ |
341 |
||
342 |
lemma ball_simps3: |
|
46820 | 343 |
"(\<forall>x\<in>Collect(A,Q).P(x)) <-> (\<forall>x\<in>A. Q(x) \<longrightarrow> P(x))" |
13780 | 344 |
by blast+ |
345 |
||
346 |
lemmas ball_simps [simp] = ball_simps1 ball_simps2 ball_simps3 |
|
347 |
||
348 |
lemma ball_conj_distrib: |
|
46820 | 349 |
"(\<forall>x\<in>A. P(x) & Q(x)) <-> ((\<forall>x\<in>A. P(x)) & (\<forall>x\<in>A. Q(x)))" |
13780 | 350 |
by blast |
351 |
||
352 |
||
353 |
subsection{*Miniscoping of the Bounded Existential Quantifier*} |
|
354 |
||
355 |
lemma bex_simps1: |
|
46820 | 356 |
"(\<exists>x\<in>A. P(x) & Q) <-> ((\<exists>x\<in>A. P(x)) & Q)" |
357 |
"(\<exists>x\<in>A. P(x) | Q) <-> (\<exists>x\<in>A. P(x)) | (A\<noteq>0 & Q)" |
|
358 |
"(\<exists>x\<in>A. P(x) \<longrightarrow> Q) <-> ((\<forall>x\<in>A. P(x)) \<longrightarrow> (A\<noteq>0 & Q))" |
|
359 |
"(\<exists>x\<in>0.P(x)) <-> False" |
|
360 |
"(\<exists>x\<in>succ(i).P(x)) <-> P(i) | (\<exists>x\<in>i. P(x))" |
|
361 |
"(\<exists>x\<in>cons(a,B).P(x)) <-> P(a) | (\<exists>x\<in>B. P(x))" |
|
362 |
"(\<exists>x\<in>RepFun(A,f). P(x)) <-> (\<exists>y\<in>A. P(f(y)))" |
|
363 |
"(\<exists>x\<in>\<Union>(A).P(x)) <-> (\<exists>y\<in>A. \<exists>x\<in>y. P(x))" |
|
364 |
"(~(\<exists>x\<in>A. P(x))) <-> (\<forall>x\<in>A. ~P(x))" |
|
13780 | 365 |
by blast+ |
366 |
||
367 |
lemma bex_simps2: |
|
46820 | 368 |
"(\<exists>x\<in>A. P & Q(x)) <-> (P & (\<exists>x\<in>A. Q(x)))" |
369 |
"(\<exists>x\<in>A. P | Q(x)) <-> (A\<noteq>0 & P) | (\<exists>x\<in>A. Q(x))" |
|
370 |
"(\<exists>x\<in>A. P \<longrightarrow> Q(x)) <-> ((A=0 | P) \<longrightarrow> (\<exists>x\<in>A. Q(x)))" |
|
13780 | 371 |
by blast+ |
372 |
||
373 |
lemma bex_simps3: |
|
46820 | 374 |
"(\<exists>x\<in>Collect(A,Q).P(x)) <-> (\<exists>x\<in>A. Q(x) & P(x))" |
13780 | 375 |
by blast |
376 |
||
377 |
lemmas bex_simps [simp] = bex_simps1 bex_simps2 bex_simps3 |
|
378 |
||
379 |
lemma bex_disj_distrib: |
|
46820 | 380 |
"(\<exists>x\<in>A. P(x) | Q(x)) <-> ((\<exists>x\<in>A. P(x)) | (\<exists>x\<in>A. Q(x)))" |
13780 | 381 |
by blast |
382 |
||
383 |
||
384 |
(** One-point rule for bounded quantifiers: see HOL/Set.ML **) |
|
385 |
||
46820 | 386 |
lemma bex_triv_one_point1 [simp]: "(\<exists>x\<in>A. x=a) <-> (a:A)" |
13780 | 387 |
by blast |
388 |
||
46820 | 389 |
lemma bex_triv_one_point2 [simp]: "(\<exists>x\<in>A. a=x) <-> (a:A)" |
13780 | 390 |
by blast |
391 |
||
46820 | 392 |
lemma bex_one_point1 [simp]: "(\<exists>x\<in>A. x=a & P(x)) <-> (a:A & P(a))" |
13780 | 393 |
by blast |
394 |
||
46820 | 395 |
lemma bex_one_point2 [simp]: "(\<exists>x\<in>A. a=x & P(x)) <-> (a:A & P(a))" |
13780 | 396 |
by blast |
397 |
||
46820 | 398 |
lemma ball_one_point1 [simp]: "(\<forall>x\<in>A. x=a \<longrightarrow> P(x)) <-> (a:A \<longrightarrow> P(a))" |
13780 | 399 |
by blast |
400 |
||
46820 | 401 |
lemma ball_one_point2 [simp]: "(\<forall>x\<in>A. a=x \<longrightarrow> P(x)) <-> (a:A \<longrightarrow> P(a))" |
13780 | 402 |
by blast |
403 |
||
404 |
||
405 |
subsection{*Miniscoping of the Replacement Operator*} |
|
406 |
||
407 |
text{*These cover both @{term Replace} and @{term Collect}*} |
|
408 |
lemma Rep_simps [simp]: |
|
409 |
"{x. y:0, R(x,y)} = 0" |
|
410 |
"{x:0. P(x)} = 0" |
|
411 |
"{x:A. Q} = (if Q then A else 0)" |
|
412 |
"RepFun(0,f) = 0" |
|
413 |
"RepFun(succ(i),f) = cons(f(i), RepFun(i,f))" |
|
414 |
"RepFun(cons(a,B),f) = cons(f(a), RepFun(B,f))" |
|
415 |
by (simp_all, blast+) |
|
416 |
||
417 |
||
418 |
subsection{*Miniscoping of Unions*} |
|
419 |
||
420 |
lemma UN_simps1: |
|
46820 | 421 |
"(\<Union>x\<in>C. cons(a, B(x))) = (if C=0 then 0 else cons(a, \<Union>x\<in>C. B(x)))" |
422 |
"(\<Union>x\<in>C. A(x) \<union> B') = (if C=0 then 0 else (\<Union>x\<in>C. A(x)) \<union> B')" |
|
423 |
"(\<Union>x\<in>C. A' \<union> B(x)) = (if C=0 then 0 else A' \<union> (\<Union>x\<in>C. B(x)))" |
|
424 |
"(\<Union>x\<in>C. A(x) \<inter> B') = ((\<Union>x\<in>C. A(x)) \<inter> B')" |
|
425 |
"(\<Union>x\<in>C. A' \<inter> B(x)) = (A' \<inter> (\<Union>x\<in>C. B(x)))" |
|
426 |
"(\<Union>x\<in>C. A(x) - B') = ((\<Union>x\<in>C. A(x)) - B')" |
|
427 |
"(\<Union>x\<in>C. A' - B(x)) = (if C=0 then 0 else A' - (\<Inter>x\<in>C. B(x)))" |
|
428 |
apply (simp_all add: Inter_def) |
|
13780 | 429 |
apply (blast intro!: equalityI )+ |
430 |
done |
|
431 |
||
432 |
lemma UN_simps2: |
|
46820 | 433 |
"(\<Union>x\<in>\<Union>(A). B(x)) = (\<Union>y\<in>A. \<Union>x\<in>y. B(x))" |
434 |
"(\<Union>z\<in>(\<Union>x\<in>A. B(x)). C(z)) = (\<Union>x\<in>A. \<Union>z\<in>B(x). C(z))" |
|
435 |
"(\<Union>x\<in>RepFun(A,f). B(x)) = (\<Union>a\<in>A. B(f(a)))" |
|
13780 | 436 |
by blast+ |
437 |
||
438 |
lemmas UN_simps [simp] = UN_simps1 UN_simps2 |
|
439 |
||
440 |
text{*Opposite of miniscoping: pull the operator out*} |
|
441 |
||
442 |
lemma UN_extend_simps1: |
|
46820 | 443 |
"(\<Union>x\<in>C. A(x)) \<union> B = (if C=0 then B else (\<Union>x\<in>C. A(x) \<union> B))" |
444 |
"((\<Union>x\<in>C. A(x)) \<inter> B) = (\<Union>x\<in>C. A(x) \<inter> B)" |
|
445 |
"((\<Union>x\<in>C. A(x)) - B) = (\<Union>x\<in>C. A(x) - B)" |
|
446 |
apply simp_all |
|
13780 | 447 |
apply blast+ |
448 |
done |
|
449 |
||
450 |
lemma UN_extend_simps2: |
|
46820 | 451 |
"cons(a, \<Union>x\<in>C. B(x)) = (if C=0 then {a} else (\<Union>x\<in>C. cons(a, B(x))))" |
452 |
"A \<union> (\<Union>x\<in>C. B(x)) = (if C=0 then A else (\<Union>x\<in>C. A \<union> B(x)))" |
|
453 |
"(A \<inter> (\<Union>x\<in>C. B(x))) = (\<Union>x\<in>C. A \<inter> B(x))" |
|
454 |
"A - (\<Inter>x\<in>C. B(x)) = (if C=0 then A else (\<Union>x\<in>C. A - B(x)))" |
|
455 |
"(\<Union>y\<in>A. \<Union>x\<in>y. B(x)) = (\<Union>x\<in>\<Union>(A). B(x))" |
|
456 |
"(\<Union>a\<in>A. B(f(a))) = (\<Union>x\<in>RepFun(A,f). B(x))" |
|
457 |
apply (simp_all add: Inter_def) |
|
13780 | 458 |
apply (blast intro!: equalityI)+ |
459 |
done |
|
460 |
||
461 |
lemma UN_UN_extend: |
|
46820 | 462 |
"(\<Union>x\<in>A. \<Union>z\<in>B(x). C(z)) = (\<Union>z\<in>(\<Union>x\<in>A. B(x)). C(z))" |
13780 | 463 |
by blast |
464 |
||
465 |
lemmas UN_extend_simps = UN_extend_simps1 UN_extend_simps2 UN_UN_extend |
|
466 |
||
467 |
||
468 |
subsection{*Miniscoping of Intersections*} |
|
469 |
||
470 |
lemma INT_simps1: |
|
46820 | 471 |
"(\<Inter>x\<in>C. A(x) \<inter> B) = (\<Inter>x\<in>C. A(x)) \<inter> B" |
472 |
"(\<Inter>x\<in>C. A(x) - B) = (\<Inter>x\<in>C. A(x)) - B" |
|
473 |
"(\<Inter>x\<in>C. A(x) \<union> B) = (if C=0 then 0 else (\<Inter>x\<in>C. A(x)) \<union> B)" |
|
13780 | 474 |
by (simp_all add: Inter_def, blast+) |
475 |
||
476 |
lemma INT_simps2: |
|
46820 | 477 |
"(\<Inter>x\<in>C. A \<inter> B(x)) = A \<inter> (\<Inter>x\<in>C. B(x))" |
478 |
"(\<Inter>x\<in>C. A - B(x)) = (if C=0 then 0 else A - (\<Union>x\<in>C. B(x)))" |
|
479 |
"(\<Inter>x\<in>C. cons(a, B(x))) = (if C=0 then 0 else cons(a, \<Inter>x\<in>C. B(x)))" |
|
480 |
"(\<Inter>x\<in>C. A \<union> B(x)) = (if C=0 then 0 else A \<union> (\<Inter>x\<in>C. B(x)))" |
|
481 |
apply (simp_all add: Inter_def) |
|
13780 | 482 |
apply (blast intro!: equalityI)+ |
483 |
done |
|
484 |
||
485 |
lemmas INT_simps [simp] = INT_simps1 INT_simps2 |
|
486 |
||
487 |
text{*Opposite of miniscoping: pull the operator out*} |
|
488 |
||
489 |
||
490 |
lemma INT_extend_simps1: |
|
46820 | 491 |
"(\<Inter>x\<in>C. A(x)) \<inter> B = (\<Inter>x\<in>C. A(x) \<inter> B)" |
492 |
"(\<Inter>x\<in>C. A(x)) - B = (\<Inter>x\<in>C. A(x) - B)" |
|
493 |
"(\<Inter>x\<in>C. A(x)) \<union> B = (if C=0 then B else (\<Inter>x\<in>C. A(x) \<union> B))" |
|
13780 | 494 |
apply (simp_all add: Inter_def, blast+) |
495 |
done |
|
496 |
||
497 |
lemma INT_extend_simps2: |
|
46820 | 498 |
"A \<inter> (\<Inter>x\<in>C. B(x)) = (\<Inter>x\<in>C. A \<inter> B(x))" |
499 |
"A - (\<Union>x\<in>C. B(x)) = (if C=0 then A else (\<Inter>x\<in>C. A - B(x)))" |
|
500 |
"cons(a, \<Inter>x\<in>C. B(x)) = (if C=0 then {a} else (\<Inter>x\<in>C. cons(a, B(x))))" |
|
501 |
"A \<union> (\<Inter>x\<in>C. B(x)) = (if C=0 then A else (\<Inter>x\<in>C. A \<union> B(x)))" |
|
502 |
apply (simp_all add: Inter_def) |
|
13780 | 503 |
apply (blast intro!: equalityI)+ |
504 |
done |
|
505 |
||
506 |
lemmas INT_extend_simps = INT_extend_simps1 INT_extend_simps2 |
|
507 |
||
508 |
||
509 |
subsection{*Other simprules*} |
|
510 |
||
511 |
||
512 |
(*** Miniscoping: pushing in big Unions, Intersections, quantifiers, etc. ***) |
|
513 |
||
514 |
lemma misc_simps [simp]: |
|
46820 | 515 |
"0 \<union> A = A" |
516 |
"A \<union> 0 = A" |
|
517 |
"0 \<inter> A = 0" |
|
518 |
"A \<inter> 0 = 0" |
|
13780 | 519 |
"0 - A = 0" |
520 |
"A - 0 = A" |
|
46820 | 521 |
"\<Union>(0) = 0" |
522 |
"\<Union>(cons(b,A)) = b \<union> \<Union>(A)" |
|
523 |
"\<Inter>({b}) = b" |
|
13780 | 524 |
by blast+ |
525 |
||
6153 | 526 |
end |