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