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