| author | wenzelm | 
| Tue, 27 Aug 2002 11:09:33 +0200 | |
| changeset 13534 | ca6debb89d77 | 
| parent 13396 | 11219ca224ab | 
| child 13544 | 895994073bdf | 
| permissions | -rw-r--r-- | 
| 1478 | 1  | 
(* Title: ZF/Ordinal.thy  | 
| 435 | 2  | 
ID: $Id$  | 
| 1478 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 435 | 4  | 
Copyright 1994 University of Cambridge  | 
5  | 
||
6  | 
*)  | 
|
7  | 
||
| 13356 | 8  | 
header{*Transitive Sets and Ordinals*}
 | 
9  | 
||
| 13155 | 10  | 
theory Ordinal = WF + Bool + equalities:  | 
11  | 
||
12  | 
constdefs  | 
|
13  | 
||
14  | 
Memrel :: "i=>i"  | 
|
15  | 
    "Memrel(A)   == {z: A*A . EX x y. z=<x,y> & x:y }"
 | 
|
16  | 
||
17  | 
Transset :: "i=>o"  | 
|
18  | 
"Transset(i) == ALL x:i. x<=i"  | 
|
19  | 
||
20  | 
Ord :: "i=>o"  | 
|
21  | 
"Ord(i) == Transset(i) & (ALL x:i. Transset(x))"  | 
|
22  | 
||
23  | 
lt :: "[i,i] => o" (infixl "<" 50) (*less-than on ordinals*)  | 
|
24  | 
"i<j == i:j & Ord(j)"  | 
|
25  | 
||
26  | 
Limit :: "i=>o"  | 
|
27  | 
"Limit(i) == Ord(i) & 0<i & (ALL y. y<i --> succ(y)<i)"  | 
|
| 2539 | 28  | 
|
29  | 
syntax  | 
|
| 13155 | 30  | 
"le" :: "[i,i] => o" (infixl 50) (*less-than or equals*)  | 
| 435 | 31  | 
|
32  | 
translations  | 
|
33  | 
"x le y" == "x < succ(y)"  | 
|
34  | 
||
| 
12114
 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 
wenzelm 
parents: 
2540 
diff
changeset
 | 
35  | 
syntax (xsymbols)  | 
| 13155 | 36  | 
"op le" :: "[i,i] => o" (infixl "\<le>" 50) (*less-than or equals*)  | 
37  | 
||
38  | 
||
| 13356 | 39  | 
subsection{*Rules for Transset*}
 | 
| 13155 | 40  | 
|
| 13356 | 41  | 
subsubsection{*Three Neat Characterisations of Transset*}
 | 
| 13155 | 42  | 
|
43  | 
lemma Transset_iff_Pow: "Transset(A) <-> A<=Pow(A)"  | 
|
44  | 
by (unfold Transset_def, blast)  | 
|
45  | 
||
46  | 
lemma Transset_iff_Union_succ: "Transset(A) <-> Union(succ(A)) = A"  | 
|
47  | 
apply (unfold Transset_def)  | 
|
48  | 
apply (blast elim!: equalityE)  | 
|
49  | 
done  | 
|
50  | 
||
51  | 
lemma Transset_iff_Union_subset: "Transset(A) <-> Union(A) <= A"  | 
|
52  | 
by (unfold Transset_def, blast)  | 
|
53  | 
||
| 13356 | 54  | 
subsubsection{*Consequences of Downwards Closure*}
 | 
| 13155 | 55  | 
|
56  | 
lemma Transset_doubleton_D:  | 
|
57  | 
    "[| Transset(C); {a,b}: C |] ==> a:C & b: C"
 | 
|
58  | 
by (unfold Transset_def, blast)  | 
|
59  | 
||
60  | 
lemma Transset_Pair_D:  | 
|
61  | 
"[| Transset(C); <a,b>: C |] ==> a:C & b: C"  | 
|
62  | 
apply (simp add: Pair_def)  | 
|
63  | 
apply (blast dest: Transset_doubleton_D)  | 
|
64  | 
done  | 
|
65  | 
||
66  | 
lemma Transset_includes_domain:  | 
|
67  | 
"[| Transset(C); A*B <= C; b: B |] ==> A <= C"  | 
|
68  | 
by (blast dest: Transset_Pair_D)  | 
|
69  | 
||
70  | 
lemma Transset_includes_range:  | 
|
71  | 
"[| Transset(C); A*B <= C; a: A |] ==> B <= C"  | 
|
72  | 
by (blast dest: Transset_Pair_D)  | 
|
73  | 
||
| 13356 | 74  | 
subsubsection{*Closure Properties*}
 | 
| 13155 | 75  | 
|
76  | 
lemma Transset_0: "Transset(0)"  | 
|
77  | 
by (unfold Transset_def, blast)  | 
|
78  | 
||
79  | 
lemma Transset_Un:  | 
|
80  | 
"[| Transset(i); Transset(j) |] ==> Transset(i Un j)"  | 
|
81  | 
by (unfold Transset_def, blast)  | 
|
82  | 
||
83  | 
lemma Transset_Int:  | 
|
84  | 
"[| Transset(i); Transset(j) |] ==> Transset(i Int j)"  | 
|
85  | 
by (unfold Transset_def, blast)  | 
|
86  | 
||
87  | 
lemma Transset_succ: "Transset(i) ==> Transset(succ(i))"  | 
|
88  | 
by (unfold Transset_def, blast)  | 
|
89  | 
||
90  | 
lemma Transset_Pow: "Transset(i) ==> Transset(Pow(i))"  | 
|
91  | 
by (unfold Transset_def, blast)  | 
|
92  | 
||
93  | 
lemma Transset_Union: "Transset(A) ==> Transset(Union(A))"  | 
|
94  | 
by (unfold Transset_def, blast)  | 
|
95  | 
||
96  | 
lemma Transset_Union_family:  | 
|
97  | 
"[| !!i. i:A ==> Transset(i) |] ==> Transset(Union(A))"  | 
|
98  | 
by (unfold Transset_def, blast)  | 
|
99  | 
||
100  | 
lemma Transset_Inter_family:  | 
|
| 
13203
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
101  | 
"[| !!i. i:A ==> Transset(i) |] ==> Transset(Inter(A))"  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
102  | 
by (unfold Inter_def Transset_def, blast)  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
103  | 
|
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
104  | 
lemma Transset_UN:  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
105  | 
"(!!x. x \<in> A ==> Transset(B(x))) ==> Transset (UN x:A. B(x))"  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
106  | 
by (rule Transset_Union_family, auto)  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
107  | 
|
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
108  | 
lemma Transset_INT:  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
109  | 
"(!!x. x \<in> A ==> Transset(B(x))) ==> Transset (INT x:A. B(x))"  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
110  | 
by (rule Transset_Inter_family, auto)  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
111  | 
|
| 13155 | 112  | 
|
| 13356 | 113  | 
subsection{*Lemmas for Ordinals*}
 | 
| 13155 | 114  | 
|
115  | 
lemma OrdI:  | 
|
116  | 
"[| Transset(i); !!x. x:i ==> Transset(x) |] ==> Ord(i)"  | 
|
117  | 
by (simp add: Ord_def)  | 
|
118  | 
||
119  | 
lemma Ord_is_Transset: "Ord(i) ==> Transset(i)"  | 
|
120  | 
by (simp add: Ord_def)  | 
|
121  | 
||
122  | 
lemma Ord_contains_Transset:  | 
|
123  | 
"[| Ord(i); j:i |] ==> Transset(j) "  | 
|
124  | 
by (unfold Ord_def, blast)  | 
|
125  | 
||
126  | 
||
127  | 
lemma Ord_in_Ord: "[| Ord(i); j:i |] ==> Ord(j)"  | 
|
128  | 
by (unfold Ord_def Transset_def, blast)  | 
|
129  | 
||
| 13243 | 130  | 
(*suitable for rewriting PROVIDED i has been fixed*)  | 
131  | 
lemma Ord_in_Ord': "[| j:i; Ord(i) |] ==> Ord(j)"  | 
|
132  | 
by (blast intro: Ord_in_Ord)  | 
|
133  | 
||
| 13155 | 134  | 
(* Ord(succ(j)) ==> Ord(j) *)  | 
135  | 
lemmas Ord_succD = Ord_in_Ord [OF _ succI1]  | 
|
136  | 
||
137  | 
lemma Ord_subset_Ord: "[| Ord(i); Transset(j); j<=i |] ==> Ord(j)"  | 
|
138  | 
by (simp add: Ord_def Transset_def, blast)  | 
|
139  | 
||
140  | 
lemma OrdmemD: "[| j:i; Ord(i) |] ==> j<=i"  | 
|
141  | 
by (unfold Ord_def Transset_def, blast)  | 
|
142  | 
||
143  | 
lemma Ord_trans: "[| i:j; j:k; Ord(k) |] ==> i:k"  | 
|
144  | 
by (blast dest: OrdmemD)  | 
|
145  | 
||
146  | 
lemma Ord_succ_subsetI: "[| i:j; Ord(j) |] ==> succ(i) <= j"  | 
|
147  | 
by (blast dest: OrdmemD)  | 
|
148  | 
||
149  | 
||
| 13356 | 150  | 
subsection{*The Construction of Ordinals: 0, succ, Union*}
 | 
| 13155 | 151  | 
|
152  | 
lemma Ord_0 [iff,TC]: "Ord(0)"  | 
|
153  | 
by (blast intro: OrdI Transset_0)  | 
|
154  | 
||
155  | 
lemma Ord_succ [TC]: "Ord(i) ==> Ord(succ(i))"  | 
|
156  | 
by (blast intro: OrdI Transset_succ Ord_is_Transset Ord_contains_Transset)  | 
|
157  | 
||
158  | 
lemmas Ord_1 = Ord_0 [THEN Ord_succ]  | 
|
159  | 
||
160  | 
lemma Ord_succ_iff [iff]: "Ord(succ(i)) <-> Ord(i)"  | 
|
161  | 
by (blast intro: Ord_succ dest!: Ord_succD)  | 
|
162  | 
||
| 13172 | 163  | 
lemma Ord_Un [intro,simp,TC]: "[| Ord(i); Ord(j) |] ==> Ord(i Un j)"  | 
| 13155 | 164  | 
apply (unfold Ord_def)  | 
165  | 
apply (blast intro!: Transset_Un)  | 
|
166  | 
done  | 
|
167  | 
||
168  | 
lemma Ord_Int [TC]: "[| Ord(i); Ord(j) |] ==> Ord(i Int j)"  | 
|
169  | 
apply (unfold Ord_def)  | 
|
170  | 
apply (blast intro!: Transset_Int)  | 
|
171  | 
done  | 
|
172  | 
||
173  | 
(*There is no set of all ordinals, for then it would contain itself*)  | 
|
174  | 
lemma ON_class: "~ (ALL i. i:X <-> Ord(i))"  | 
|
175  | 
apply (rule notI)  | 
|
176  | 
apply (frule_tac x = "X" in spec)  | 
|
177  | 
apply (safe elim!: mem_irrefl)  | 
|
178  | 
apply (erule swap, rule OrdI [OF _ Ord_is_Transset])  | 
|
179  | 
apply (simp add: Transset_def)  | 
|
180  | 
apply (blast intro: Ord_in_Ord)+  | 
|
181  | 
done  | 
|
182  | 
||
| 13356 | 183  | 
subsection{*< is 'less Than' for Ordinals*}
 | 
| 13155 | 184  | 
|
185  | 
lemma ltI: "[| i:j; Ord(j) |] ==> i<j"  | 
|
186  | 
by (unfold lt_def, blast)  | 
|
187  | 
||
188  | 
lemma ltE:  | 
|
189  | 
"[| i<j; [| i:j; Ord(i); Ord(j) |] ==> P |] ==> P"  | 
|
190  | 
apply (unfold lt_def)  | 
|
191  | 
apply (blast intro: Ord_in_Ord)  | 
|
192  | 
done  | 
|
193  | 
||
194  | 
lemma ltD: "i<j ==> i:j"  | 
|
195  | 
by (erule ltE, assumption)  | 
|
196  | 
||
197  | 
lemma not_lt0 [simp]: "~ i<0"  | 
|
198  | 
by (unfold lt_def, blast)  | 
|
199  | 
||
200  | 
lemma lt_Ord: "j<i ==> Ord(j)"  | 
|
201  | 
by (erule ltE, assumption)  | 
|
202  | 
||
203  | 
lemma lt_Ord2: "j<i ==> Ord(i)"  | 
|
204  | 
by (erule ltE, assumption)  | 
|
205  | 
||
206  | 
(* "ja le j ==> Ord(j)" *)  | 
|
207  | 
lemmas le_Ord2 = lt_Ord2 [THEN Ord_succD]  | 
|
208  | 
||
209  | 
(* i<0 ==> R *)  | 
|
210  | 
lemmas lt0E = not_lt0 [THEN notE, elim!]  | 
|
211  | 
||
212  | 
lemma lt_trans: "[| i<j; j<k |] ==> i<k"  | 
|
213  | 
by (blast intro!: ltI elim!: ltE intro: Ord_trans)  | 
|
214  | 
||
215  | 
lemma lt_not_sym: "i<j ==> ~ (j<i)"  | 
|
216  | 
apply (unfold lt_def)  | 
|
217  | 
apply (blast elim: mem_asym)  | 
|
218  | 
done  | 
|
219  | 
||
220  | 
(* [| i<j; ~P ==> j<i |] ==> P *)  | 
|
221  | 
lemmas lt_asym = lt_not_sym [THEN swap]  | 
|
222  | 
||
223  | 
lemma lt_irrefl [elim!]: "i<i ==> P"  | 
|
224  | 
by (blast intro: lt_asym)  | 
|
225  | 
||
226  | 
lemma lt_not_refl: "~ i<i"  | 
|
227  | 
apply (rule notI)  | 
|
228  | 
apply (erule lt_irrefl)  | 
|
229  | 
done  | 
|
230  | 
||
231  | 
||
232  | 
(** le is less than or equals; recall i le j abbrevs i<succ(j) !! **)  | 
|
233  | 
||
234  | 
lemma le_iff: "i le j <-> i<j | (i=j & Ord(j))"  | 
|
235  | 
by (unfold lt_def, blast)  | 
|
236  | 
||
237  | 
(*Equivalently, i<j ==> i < succ(j)*)  | 
|
238  | 
lemma leI: "i<j ==> i le j"  | 
|
239  | 
by (simp (no_asm_simp) add: le_iff)  | 
|
240  | 
||
241  | 
lemma le_eqI: "[| i=j; Ord(j) |] ==> i le j"  | 
|
242  | 
by (simp (no_asm_simp) add: le_iff)  | 
|
243  | 
||
244  | 
lemmas le_refl = refl [THEN le_eqI]  | 
|
245  | 
||
246  | 
lemma le_refl_iff [iff]: "i le i <-> Ord(i)"  | 
|
247  | 
by (simp (no_asm_simp) add: lt_not_refl le_iff)  | 
|
248  | 
||
249  | 
lemma leCI: "(~ (i=j & Ord(j)) ==> i<j) ==> i le j"  | 
|
250  | 
by (simp add: le_iff, blast)  | 
|
251  | 
||
252  | 
lemma leE:  | 
|
253  | 
"[| i le j; i<j ==> P; [| i=j; Ord(j) |] ==> P |] ==> P"  | 
|
254  | 
by (simp add: le_iff, blast)  | 
|
255  | 
||
256  | 
lemma le_anti_sym: "[| i le j; j le i |] ==> i=j"  | 
|
257  | 
apply (simp add: le_iff)  | 
|
258  | 
apply (blast elim: lt_asym)  | 
|
259  | 
done  | 
|
260  | 
||
261  | 
lemma le0_iff [simp]: "i le 0 <-> i=0"  | 
|
262  | 
by (blast elim!: leE)  | 
|
263  | 
||
264  | 
lemmas le0D = le0_iff [THEN iffD1, dest!]  | 
|
265  | 
||
| 13356 | 266  | 
subsection{*Natural Deduction Rules for Memrel*}
 | 
| 13155 | 267  | 
|
268  | 
(*The lemmas MemrelI/E give better speed than [iff] here*)  | 
|
269  | 
lemma Memrel_iff [simp]: "<a,b> : Memrel(A) <-> a:b & a:A & b:A"  | 
|
270  | 
by (unfold Memrel_def, blast)  | 
|
271  | 
||
272  | 
lemma MemrelI [intro!]: "[| a: b; a: A; b: A |] ==> <a,b> : Memrel(A)"  | 
|
273  | 
by auto  | 
|
274  | 
||
275  | 
lemma MemrelE [elim!]:  | 
|
276  | 
"[| <a,b> : Memrel(A);  | 
|
277  | 
[| a: A; b: A; a:b |] ==> P |]  | 
|
278  | 
==> P"  | 
|
279  | 
by auto  | 
|
280  | 
||
281  | 
lemma Memrel_type: "Memrel(A) <= A*A"  | 
|
282  | 
by (unfold Memrel_def, blast)  | 
|
283  | 
||
284  | 
lemma Memrel_mono: "A<=B ==> Memrel(A) <= Memrel(B)"  | 
|
285  | 
by (unfold Memrel_def, blast)  | 
|
286  | 
||
287  | 
lemma Memrel_0 [simp]: "Memrel(0) = 0"  | 
|
288  | 
by (unfold Memrel_def, blast)  | 
|
289  | 
||
290  | 
lemma Memrel_1 [simp]: "Memrel(1) = 0"  | 
|
291  | 
by (unfold Memrel_def, blast)  | 
|
292  | 
||
| 13269 | 293  | 
lemma relation_Memrel: "relation(Memrel(A))"  | 
294  | 
by (simp add: relation_def Memrel_def, blast)  | 
|
295  | 
||
| 13155 | 296  | 
(*The membership relation (as a set) is well-founded.  | 
297  | 
Proof idea: show A<=B by applying the foundation axiom to A-B *)  | 
|
298  | 
lemma wf_Memrel: "wf(Memrel(A))"  | 
|
299  | 
apply (unfold wf_def)  | 
|
300  | 
apply (rule foundation [THEN disjE, THEN allI], erule disjI1, blast)  | 
|
301  | 
done  | 
|
302  | 
||
| 13396 | 303  | 
text{*The premise @{term "Ord(i)"} does not suffice.*}
 | 
| 13155 | 304  | 
lemma trans_Memrel:  | 
305  | 
"Ord(i) ==> trans(Memrel(i))"  | 
|
306  | 
by (unfold Ord_def Transset_def trans_def, blast)  | 
|
307  | 
||
| 13396 | 308  | 
text{*However, the following premise is strong enough.*}
 | 
309  | 
lemma Transset_trans_Memrel:  | 
|
310  | 
"\<forall>j\<in>i. Transset(j) ==> trans(Memrel(i))"  | 
|
311  | 
by (unfold Transset_def trans_def, blast)  | 
|
312  | 
||
| 13155 | 313  | 
(*If Transset(A) then Memrel(A) internalizes the membership relation below A*)  | 
314  | 
lemma Transset_Memrel_iff:  | 
|
315  | 
"Transset(A) ==> <a,b> : Memrel(A) <-> a:b & b:A"  | 
|
316  | 
by (unfold Transset_def, blast)  | 
|
317  | 
||
318  | 
||
| 13356 | 319  | 
subsection{*Transfinite Induction*}
 | 
| 13155 | 320  | 
|
321  | 
(*Epsilon induction over a transitive set*)  | 
|
322  | 
lemma Transset_induct:  | 
|
323  | 
"[| i: k; Transset(k);  | 
|
324  | 
!!x.[| x: k; ALL y:x. P(y) |] ==> P(x) |]  | 
|
325  | 
==> P(i)"  | 
|
326  | 
apply (simp add: Transset_def)  | 
|
| 13269 | 327  | 
apply (erule wf_Memrel [THEN wf_induct2], blast+)  | 
| 13155 | 328  | 
done  | 
329  | 
||
330  | 
(*Induction over an ordinal*)  | 
|
| 13534 | 331  | 
lemmas Ord_induct [consumes 2] = Transset_induct [OF _ Ord_is_Transset]  | 
332  | 
lemmas Ord_induct_rule = Ord_induct [rule_format, consumes 2]  | 
|
| 13155 | 333  | 
|
334  | 
(*Induction over the class of ordinals -- a useful corollary of Ord_induct*)  | 
|
335  | 
||
| 13534 | 336  | 
lemma trans_induct [consumes 1]:  | 
| 13155 | 337  | 
"[| Ord(i);  | 
338  | 
!!x.[| Ord(x); ALL y:x. P(y) |] ==> P(x) |]  | 
|
339  | 
==> P(i)"  | 
|
340  | 
apply (rule Ord_succ [THEN succI1 [THEN Ord_induct]], assumption)  | 
|
341  | 
apply (blast intro: Ord_succ [THEN Ord_in_Ord])  | 
|
342  | 
done  | 
|
343  | 
||
| 13534 | 344  | 
lemmas trans_induct_rule = trans_induct [rule_format, consumes 1]  | 
345  | 
||
| 13155 | 346  | 
|
347  | 
(*** Fundamental properties of the epsilon ordering (< on ordinals) ***)  | 
|
348  | 
||
349  | 
||
| 13356 | 350  | 
subsubsection{*Proving That < is a Linear Ordering on the Ordinals*}
 | 
| 13155 | 351  | 
|
352  | 
lemma Ord_linear [rule_format]:  | 
|
353  | 
"Ord(i) ==> (ALL j. Ord(j) --> i:j | i=j | j:i)"  | 
|
354  | 
apply (erule trans_induct)  | 
|
355  | 
apply (rule impI [THEN allI])  | 
|
356  | 
apply (erule_tac i=j in trans_induct)  | 
|
357  | 
apply (blast dest: Ord_trans)  | 
|
358  | 
done  | 
|
359  | 
||
360  | 
(*The trichotomy law for ordinals!*)  | 
|
361  | 
lemma Ord_linear_lt:  | 
|
362  | 
"[| Ord(i); Ord(j); i<j ==> P; i=j ==> P; j<i ==> P |] ==> P"  | 
|
363  | 
apply (simp add: lt_def)  | 
|
364  | 
apply (rule_tac i1=i and j1=j in Ord_linear [THEN disjE], blast+)  | 
|
365  | 
done  | 
|
366  | 
||
367  | 
lemma Ord_linear2:  | 
|
368  | 
"[| Ord(i); Ord(j); i<j ==> P; j le i ==> P |] ==> P"  | 
|
369  | 
apply (rule_tac i = "i" and j = "j" in Ord_linear_lt)  | 
|
370  | 
apply (blast intro: leI le_eqI sym ) +  | 
|
371  | 
done  | 
|
372  | 
||
373  | 
lemma Ord_linear_le:  | 
|
374  | 
"[| Ord(i); Ord(j); i le j ==> P; j le i ==> P |] ==> P"  | 
|
375  | 
apply (rule_tac i = "i" and j = "j" in Ord_linear_lt)  | 
|
376  | 
apply (blast intro: leI le_eqI ) +  | 
|
377  | 
done  | 
|
378  | 
||
379  | 
lemma le_imp_not_lt: "j le i ==> ~ i<j"  | 
|
380  | 
by (blast elim!: leE elim: lt_asym)  | 
|
381  | 
||
382  | 
lemma not_lt_imp_le: "[| ~ i<j; Ord(i); Ord(j) |] ==> j le i"  | 
|
383  | 
by (rule_tac i = "i" and j = "j" in Ord_linear2, auto)  | 
|
384  | 
||
| 13356 | 385  | 
subsubsection{*Some Rewrite Rules for <, le*}
 | 
| 13155 | 386  | 
|
387  | 
lemma Ord_mem_iff_lt: "Ord(j) ==> i:j <-> i<j"  | 
|
388  | 
by (unfold lt_def, blast)  | 
|
389  | 
||
390  | 
lemma not_lt_iff_le: "[| Ord(i); Ord(j) |] ==> ~ i<j <-> j le i"  | 
|
391  | 
by (blast dest: le_imp_not_lt not_lt_imp_le)  | 
|
| 2540 | 392  | 
|
| 13155 | 393  | 
lemma not_le_iff_lt: "[| Ord(i); Ord(j) |] ==> ~ i le j <-> j<i"  | 
394  | 
by (simp (no_asm_simp) add: not_lt_iff_le [THEN iff_sym])  | 
|
395  | 
||
396  | 
(*This is identical to 0<succ(i) *)  | 
|
397  | 
lemma Ord_0_le: "Ord(i) ==> 0 le i"  | 
|
398  | 
by (erule not_lt_iff_le [THEN iffD1], auto)  | 
|
399  | 
||
400  | 
lemma Ord_0_lt: "[| Ord(i); i~=0 |] ==> 0<i"  | 
|
401  | 
apply (erule not_le_iff_lt [THEN iffD1])  | 
|
402  | 
apply (rule Ord_0, blast)  | 
|
403  | 
done  | 
|
404  | 
||
405  | 
lemma Ord_0_lt_iff: "Ord(i) ==> i~=0 <-> 0<i"  | 
|
406  | 
by (blast intro: Ord_0_lt)  | 
|
407  | 
||
408  | 
||
| 13356 | 409  | 
subsection{*Results about Less-Than or Equals*}
 | 
| 13155 | 410  | 
|
411  | 
(** For ordinals, j<=i (subset) implies j le i (less-than or equals) **)  | 
|
412  | 
||
413  | 
lemma zero_le_succ_iff [iff]: "0 le succ(x) <-> Ord(x)"  | 
|
414  | 
by (blast intro: Ord_0_le elim: ltE)  | 
|
415  | 
||
416  | 
lemma subset_imp_le: "[| j<=i; Ord(i); Ord(j) |] ==> j le i"  | 
|
| 13269 | 417  | 
apply (rule not_lt_iff_le [THEN iffD1], assumption+)  | 
| 13155 | 418  | 
apply (blast elim: ltE mem_irrefl)  | 
419  | 
done  | 
|
420  | 
||
421  | 
lemma le_imp_subset: "i le j ==> i<=j"  | 
|
422  | 
by (blast dest: OrdmemD elim: ltE leE)  | 
|
423  | 
||
424  | 
lemma le_subset_iff: "j le i <-> j<=i & Ord(i) & Ord(j)"  | 
|
425  | 
by (blast dest: subset_imp_le le_imp_subset elim: ltE)  | 
|
426  | 
||
427  | 
lemma le_succ_iff: "i le succ(j) <-> i le j | i=succ(j) & Ord(i)"  | 
|
428  | 
apply (simp (no_asm) add: le_iff)  | 
|
429  | 
apply blast  | 
|
430  | 
done  | 
|
431  | 
||
432  | 
(*Just a variant of subset_imp_le*)  | 
|
433  | 
lemma all_lt_imp_le: "[| Ord(i); Ord(j); !!x. x<j ==> x<i |] ==> j le i"  | 
|
434  | 
by (blast intro: not_lt_imp_le dest: lt_irrefl)  | 
|
435  | 
||
| 13356 | 436  | 
subsubsection{*Transitivity Laws*}
 | 
| 13155 | 437  | 
|
438  | 
lemma lt_trans1: "[| i le j; j<k |] ==> i<k"  | 
|
439  | 
by (blast elim!: leE intro: lt_trans)  | 
|
440  | 
||
441  | 
lemma lt_trans2: "[| i<j; j le k |] ==> i<k"  | 
|
442  | 
by (blast elim!: leE intro: lt_trans)  | 
|
443  | 
||
444  | 
lemma le_trans: "[| i le j; j le k |] ==> i le k"  | 
|
445  | 
by (blast intro: lt_trans1)  | 
|
446  | 
||
447  | 
lemma succ_leI: "i<j ==> succ(i) le j"  | 
|
448  | 
apply (rule not_lt_iff_le [THEN iffD1])  | 
|
449  | 
apply (blast elim: ltE leE lt_asym)+  | 
|
450  | 
done  | 
|
451  | 
||
452  | 
(*Identical to succ(i) < succ(j) ==> i<j *)  | 
|
453  | 
lemma succ_leE: "succ(i) le j ==> i<j"  | 
|
454  | 
apply (rule not_le_iff_lt [THEN iffD1])  | 
|
455  | 
apply (blast elim: ltE leE lt_asym)+  | 
|
456  | 
done  | 
|
457  | 
||
458  | 
lemma succ_le_iff [iff]: "succ(i) le j <-> i<j"  | 
|
459  | 
by (blast intro: succ_leI succ_leE)  | 
|
460  | 
||
461  | 
lemma succ_le_imp_le: "succ(i) le succ(j) ==> i le j"  | 
|
462  | 
by (blast dest!: succ_leE)  | 
|
463  | 
||
464  | 
lemma lt_subset_trans: "[| i <= j; j<k; Ord(i) |] ==> i<k"  | 
|
465  | 
apply (rule subset_imp_le [THEN lt_trans1])  | 
|
466  | 
apply (blast intro: elim: ltE) +  | 
|
467  | 
done  | 
|
468  | 
||
| 13172 | 469  | 
lemma lt_imp_0_lt: "j<i ==> 0<i"  | 
470  | 
by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord])  | 
|
471  | 
||
| 13243 | 472  | 
lemma succ_lt_iff: "succ(i) < j <-> i<j & succ(i) \<noteq> j"  | 
| 
13162
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
473  | 
apply auto  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
474  | 
apply (blast intro: lt_trans le_refl dest: lt_Ord)  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
475  | 
apply (frule lt_Ord)  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
476  | 
apply (rule not_le_iff_lt [THEN iffD1])  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
477  | 
apply (blast intro: lt_Ord2)  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
478  | 
apply blast  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
479  | 
apply (simp add: lt_Ord lt_Ord2 le_iff)  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
480  | 
apply (blast dest: lt_asym)  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
481  | 
done  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
482  | 
|
| 13243 | 483  | 
lemma Ord_succ_mem_iff: "Ord(j) ==> succ(i) \<in> succ(j) <-> i\<in>j"  | 
484  | 
apply (insert succ_le_iff [of i j])  | 
|
485  | 
apply (simp add: lt_def)  | 
|
486  | 
done  | 
|
487  | 
||
| 13356 | 488  | 
subsubsection{*Union and Intersection*}
 | 
| 13155 | 489  | 
|
490  | 
lemma Un_upper1_le: "[| Ord(i); Ord(j) |] ==> i le i Un j"  | 
|
491  | 
by (rule Un_upper1 [THEN subset_imp_le], auto)  | 
|
492  | 
||
493  | 
lemma Un_upper2_le: "[| Ord(i); Ord(j) |] ==> j le i Un j"  | 
|
494  | 
by (rule Un_upper2 [THEN subset_imp_le], auto)  | 
|
495  | 
||
496  | 
(*Replacing k by succ(k') yields the similar rule for le!*)  | 
|
497  | 
lemma Un_least_lt: "[| i<k; j<k |] ==> i Un j < k"  | 
|
498  | 
apply (rule_tac i = "i" and j = "j" in Ord_linear_le)  | 
|
499  | 
apply (auto simp add: Un_commute le_subset_iff subset_Un_iff lt_Ord)  | 
|
500  | 
done  | 
|
501  | 
||
502  | 
lemma Un_least_lt_iff: "[| Ord(i); Ord(j) |] ==> i Un j < k <-> i<k & j<k"  | 
|
503  | 
apply (safe intro!: Un_least_lt)  | 
|
504  | 
apply (rule_tac [2] Un_upper2_le [THEN lt_trans1])  | 
|
505  | 
apply (rule Un_upper1_le [THEN lt_trans1], auto)  | 
|
506  | 
done  | 
|
507  | 
||
508  | 
lemma Un_least_mem_iff:  | 
|
509  | 
"[| Ord(i); Ord(j); Ord(k) |] ==> i Un j : k <-> i:k & j:k"  | 
|
510  | 
apply (insert Un_least_lt_iff [of i j k])  | 
|
511  | 
apply (simp add: lt_def)  | 
|
512  | 
done  | 
|
513  | 
||
514  | 
(*Replacing k by succ(k') yields the similar rule for le!*)  | 
|
515  | 
lemma Int_greatest_lt: "[| i<k; j<k |] ==> i Int j < k"  | 
|
516  | 
apply (rule_tac i = "i" and j = "j" in Ord_linear_le)  | 
|
517  | 
apply (auto simp add: Int_commute le_subset_iff subset_Int_iff lt_Ord)  | 
|
518  | 
done  | 
|
519  | 
||
| 
13162
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
520  | 
lemma Ord_Un_if:  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
521  | 
"[| Ord(i); Ord(j) |] ==> i \<union> j = (if j<i then i else j)"  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
522  | 
by (simp add: not_lt_iff_le le_imp_subset leI  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
523  | 
subset_Un_iff [symmetric] subset_Un_iff2 [symmetric])  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
524  | 
|
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
525  | 
lemma succ_Un_distrib:  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
526  | 
"[| Ord(i); Ord(j) |] ==> succ(i \<union> j) = succ(i) \<union> succ(j)"  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
527  | 
by (simp add: Ord_Un_if lt_Ord le_Ord2)  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
528  | 
|
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
529  | 
lemma lt_Un_iff:  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
530  | 
"[| Ord(i); Ord(j) |] ==> k < i \<union> j <-> k < i | k < j";  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
531  | 
apply (simp add: Ord_Un_if not_lt_iff_le)  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
532  | 
apply (blast intro: leI lt_trans2)+  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
533  | 
done  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
534  | 
|
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
535  | 
lemma le_Un_iff:  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
536  | 
"[| Ord(i); Ord(j) |] ==> k \<le> i \<union> j <-> k \<le> i | k \<le> j";  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
537  | 
by (simp add: succ_Un_distrib lt_Un_iff [symmetric])  | 
| 
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
538  | 
|
| 13172 | 539  | 
lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i Un j"  | 
540  | 
by (simp add: lt_Un_iff lt_Ord2)  | 
|
541  | 
||
542  | 
lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i Un j"  | 
|
543  | 
by (simp add: lt_Un_iff lt_Ord2)  | 
|
544  | 
||
545  | 
(*See also Transset_iff_Union_succ*)  | 
|
546  | 
lemma Ord_Union_succ_eq: "Ord(i) ==> \<Union>(succ(i)) = i"  | 
|
547  | 
by (blast intro: Ord_trans)  | 
|
548  | 
||
| 
13162
 
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
 
paulson 
parents: 
13155 
diff
changeset
 | 
549  | 
|
| 13356 | 550  | 
subsection{*Results about Limits*}
 | 
| 13155 | 551  | 
|
| 13172 | 552  | 
lemma Ord_Union [intro,simp,TC]: "[| !!i. i:A ==> Ord(i) |] ==> Ord(Union(A))"  | 
| 13155 | 553  | 
apply (rule Ord_is_Transset [THEN Transset_Union_family, THEN OrdI])  | 
554  | 
apply (blast intro: Ord_contains_Transset)+  | 
|
555  | 
done  | 
|
556  | 
||
| 13172 | 557  | 
lemma Ord_UN [intro,simp,TC]:  | 
558  | 
"[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(UN x:A. B(x))"  | 
|
| 13155 | 559  | 
by (rule Ord_Union, blast)  | 
560  | 
||
| 
13203
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
561  | 
lemma Ord_Inter [intro,simp,TC]:  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
562  | 
"[| !!i. i:A ==> Ord(i) |] ==> Ord(Inter(A))"  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
563  | 
apply (rule Transset_Inter_family [THEN OrdI])  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
564  | 
apply (blast intro: Ord_is_Transset)  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
565  | 
apply (simp add: Inter_def)  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
566  | 
apply (blast intro: Ord_contains_Transset)  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
567  | 
done  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
568  | 
|
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
569  | 
lemma Ord_INT [intro,simp,TC]:  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
570  | 
"[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(INT x:A. B(x))"  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
571  | 
by (rule Ord_Inter, blast)  | 
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
572  | 
|
| 
 
fac77a839aa2
Tidying up.  Mainly moving proofs from Main.thy to other (Isar) theory files.
 
paulson 
parents: 
13172 
diff
changeset
 | 
573  | 
|
| 13155 | 574  | 
(* No < version; consider (UN i:nat.i)=nat *)  | 
575  | 
lemma UN_least_le:  | 
|
576  | 
"[| Ord(i); !!x. x:A ==> b(x) le i |] ==> (UN x:A. b(x)) le i"  | 
|
577  | 
apply (rule le_imp_subset [THEN UN_least, THEN subset_imp_le])  | 
|
578  | 
apply (blast intro: Ord_UN elim: ltE)+  | 
|
579  | 
done  | 
|
580  | 
||
581  | 
lemma UN_succ_least_lt:  | 
|
582  | 
"[| j<i; !!x. x:A ==> b(x)<j |] ==> (UN x:A. succ(b(x))) < i"  | 
|
583  | 
apply (rule ltE, assumption)  | 
|
584  | 
apply (rule UN_least_le [THEN lt_trans2])  | 
|
585  | 
apply (blast intro: succ_leI)+  | 
|
586  | 
done  | 
|
587  | 
||
| 13172 | 588  | 
lemma UN_upper_lt:  | 
589  | 
"[| a\<in>A; i < b(a); Ord(\<Union>x\<in>A. b(x)) |] ==> i < (\<Union>x\<in>A. b(x))"  | 
|
590  | 
by (unfold lt_def, blast)  | 
|
591  | 
||
| 13155 | 592  | 
lemma UN_upper_le:  | 
593  | 
"[| a: A; i le b(a); Ord(UN x:A. b(x)) |] ==> i le (UN x:A. b(x))"  | 
|
594  | 
apply (frule ltD)  | 
|
595  | 
apply (rule le_imp_subset [THEN subset_trans, THEN subset_imp_le])  | 
|
596  | 
apply (blast intro: lt_Ord UN_upper)+  | 
|
597  | 
done  | 
|
598  | 
||
| 13172 | 599  | 
lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) ==> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)"  | 
600  | 
by (auto simp: lt_def Ord_Union)  | 
|
601  | 
||
602  | 
lemma Union_upper_le:  | 
|
603  | 
"[| j: J; i\<le>j; Ord(\<Union>(J)) |] ==> i \<le> \<Union>J"  | 
|
604  | 
apply (subst Union_eq_UN)  | 
|
605  | 
apply (rule UN_upper_le, auto)  | 
|
606  | 
done  | 
|
607  | 
||
| 13155 | 608  | 
lemma le_implies_UN_le_UN:  | 
609  | 
"[| !!x. x:A ==> c(x) le d(x) |] ==> (UN x:A. c(x)) le (UN x:A. d(x))"  | 
|
610  | 
apply (rule UN_least_le)  | 
|
611  | 
apply (rule_tac [2] UN_upper_le)  | 
|
612  | 
apply (blast intro: Ord_UN le_Ord2)+  | 
|
613  | 
done  | 
|
614  | 
||
615  | 
lemma Ord_equality: "Ord(i) ==> (UN y:i. succ(y)) = i"  | 
|
616  | 
by (blast intro: Ord_trans)  | 
|
617  | 
||
618  | 
(*Holds for all transitive sets, not just ordinals*)  | 
|
619  | 
lemma Ord_Union_subset: "Ord(i) ==> Union(i) <= i"  | 
|
620  | 
by (blast intro: Ord_trans)  | 
|
621  | 
||
622  | 
||
| 13356 | 623  | 
subsection{*Limit Ordinals -- General Properties*}
 | 
| 13155 | 624  | 
|
625  | 
lemma Limit_Union_eq: "Limit(i) ==> Union(i) = i"  | 
|
626  | 
apply (unfold Limit_def)  | 
|
627  | 
apply (fast intro!: ltI elim!: ltE elim: Ord_trans)  | 
|
628  | 
done  | 
|
629  | 
||
630  | 
lemma Limit_is_Ord: "Limit(i) ==> Ord(i)"  | 
|
631  | 
apply (unfold Limit_def)  | 
|
632  | 
apply (erule conjunct1)  | 
|
633  | 
done  | 
|
634  | 
||
635  | 
lemma Limit_has_0: "Limit(i) ==> 0 < i"  | 
|
636  | 
apply (unfold Limit_def)  | 
|
637  | 
apply (erule conjunct2 [THEN conjunct1])  | 
|
638  | 
done  | 
|
639  | 
||
640  | 
lemma Limit_has_succ: "[| Limit(i); j<i |] ==> succ(j) < i"  | 
|
641  | 
by (unfold Limit_def, blast)  | 
|
642  | 
||
| 13172 | 643  | 
lemma zero_not_Limit [iff]: "~ Limit(0)"  | 
644  | 
by (simp add: Limit_def)  | 
|
645  | 
||
646  | 
lemma Limit_has_1: "Limit(i) ==> 1 < i"  | 
|
647  | 
by (blast intro: Limit_has_0 Limit_has_succ)  | 
|
648  | 
||
649  | 
lemma increasing_LimitI: "[| 0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y |] ==> Limit(l)"  | 
|
650  | 
apply (simp add: Limit_def lt_Ord2, clarify)  | 
|
651  | 
apply (drule_tac i=y in ltD)  | 
|
652  | 
apply (blast intro: lt_trans1 [OF _ ltI] lt_Ord2)  | 
|
653  | 
done  | 
|
654  | 
||
| 13155 | 655  | 
lemma non_succ_LimitI:  | 
656  | 
"[| 0<i; ALL y. succ(y) ~= i |] ==> Limit(i)"  | 
|
657  | 
apply (unfold Limit_def)  | 
|
658  | 
apply (safe del: subsetI)  | 
|
659  | 
apply (rule_tac [2] not_le_iff_lt [THEN iffD1])  | 
|
660  | 
apply (simp_all add: lt_Ord lt_Ord2)  | 
|
661  | 
apply (blast elim: leE lt_asym)  | 
|
662  | 
done  | 
|
663  | 
||
664  | 
lemma succ_LimitE [elim!]: "Limit(succ(i)) ==> P"  | 
|
665  | 
apply (rule lt_irrefl)  | 
|
666  | 
apply (rule Limit_has_succ, assumption)  | 
|
667  | 
apply (erule Limit_is_Ord [THEN Ord_succD, THEN le_refl])  | 
|
668  | 
done  | 
|
669  | 
||
670  | 
lemma not_succ_Limit [simp]: "~ Limit(succ(i))"  | 
|
671  | 
by blast  | 
|
672  | 
||
673  | 
lemma Limit_le_succD: "[| Limit(i); i le succ(j) |] ==> i le j"  | 
|
674  | 
by (blast elim!: leE)  | 
|
675  | 
||
| 13172 | 676  | 
|
| 13356 | 677  | 
subsubsection{*Traditional 3-Way Case Analysis on Ordinals*}
 | 
| 13155 | 678  | 
|
679  | 
lemma Ord_cases_disj: "Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)"  | 
|
680  | 
by (blast intro!: non_succ_LimitI Ord_0_lt)  | 
|
681  | 
||
682  | 
lemma Ord_cases:  | 
|
683  | 
"[| Ord(i);  | 
|
684  | 
i=0 ==> P;  | 
|
685  | 
!!j. [| Ord(j); i=succ(j) |] ==> P;  | 
|
686  | 
Limit(i) ==> P  | 
|
687  | 
|] ==> P"  | 
|
688  | 
by (drule Ord_cases_disj, blast)  | 
|
689  | 
||
| 13534 | 690  | 
lemma trans_induct3 [case_names 0 succ limit, consumes 1]:  | 
| 13155 | 691  | 
"[| Ord(i);  | 
692  | 
P(0);  | 
|
693  | 
!!x. [| Ord(x); P(x) |] ==> P(succ(x));  | 
|
694  | 
!!x. [| Limit(x); ALL y:x. P(y) |] ==> P(x)  | 
|
695  | 
|] ==> P(i)"  | 
|
696  | 
apply (erule trans_induct)  | 
|
697  | 
apply (erule Ord_cases, blast+)  | 
|
698  | 
done  | 
|
699  | 
||
| 13534 | 700  | 
lemmas trans_induct3_rule = trans_induct3 [rule_format, case_names 0 succ limit, consumes 1]  | 
701  | 
||
| 13172 | 702  | 
text{*A set of ordinals is either empty, contains its own union, or its
 | 
703  | 
union is a limit ordinal.*}  | 
|
704  | 
lemma Ord_set_cases:  | 
|
705  | 
"\<forall>i\<in>I. Ord(i) ==> I=0 \<or> \<Union>(I) \<in> I \<or> (\<Union>(I) \<notin> I \<and> Limit(\<Union>(I)))"  | 
|
706  | 
apply (clarify elim!: not_emptyE)  | 
|
707  | 
apply (cases "\<Union>(I)" rule: Ord_cases)  | 
|
708  | 
apply (blast intro: Ord_Union)  | 
|
709  | 
apply (blast intro: subst_elem)  | 
|
710  | 
apply auto  | 
|
711  | 
apply (clarify elim!: equalityE succ_subsetE)  | 
|
712  | 
apply (simp add: Union_subset_iff)  | 
|
713  | 
apply (subgoal_tac "B = succ(j)", blast)  | 
|
714  | 
apply (rule le_anti_sym)  | 
|
715  | 
apply (simp add: le_subset_iff)  | 
|
716  | 
apply (simp add: ltI)  | 
|
717  | 
done  | 
|
718  | 
||
719  | 
text{*If the union of a set of ordinals is a successor, then it is
 | 
|
720  | 
an element of that set.*}  | 
|
721  | 
lemma Ord_Union_eq_succD: "[|\<forall>x\<in>X. Ord(x); \<Union>X = succ(j)|] ==> succ(j) \<in> X"  | 
|
722  | 
by (drule Ord_set_cases, auto)  | 
|
723  | 
||
724  | 
lemma Limit_Union [rule_format]: "[| I \<noteq> 0; \<forall>i\<in>I. Limit(i) |] ==> Limit(\<Union>I)"  | 
|
725  | 
apply (simp add: Limit_def lt_def)  | 
|
726  | 
apply (blast intro!: equalityI)  | 
|
727  | 
done  | 
|
728  | 
||
| 13155 | 729  | 
ML  | 
730  | 
{*
 | 
|
731  | 
val Memrel_def = thm "Memrel_def";  | 
|
732  | 
val Transset_def = thm "Transset_def";  | 
|
733  | 
val Ord_def = thm "Ord_def";  | 
|
734  | 
val lt_def = thm "lt_def";  | 
|
735  | 
val Limit_def = thm "Limit_def";  | 
|
736  | 
||
737  | 
val Transset_iff_Pow = thm "Transset_iff_Pow";  | 
|
738  | 
val Transset_iff_Union_succ = thm "Transset_iff_Union_succ";  | 
|
739  | 
val Transset_iff_Union_subset = thm "Transset_iff_Union_subset";  | 
|
740  | 
val Transset_doubleton_D = thm "Transset_doubleton_D";  | 
|
741  | 
val Transset_Pair_D = thm "Transset_Pair_D";  | 
|
742  | 
val Transset_includes_domain = thm "Transset_includes_domain";  | 
|
743  | 
val Transset_includes_range = thm "Transset_includes_range";  | 
|
744  | 
val Transset_0 = thm "Transset_0";  | 
|
745  | 
val Transset_Un = thm "Transset_Un";  | 
|
746  | 
val Transset_Int = thm "Transset_Int";  | 
|
747  | 
val Transset_succ = thm "Transset_succ";  | 
|
748  | 
val Transset_Pow = thm "Transset_Pow";  | 
|
749  | 
val Transset_Union = thm "Transset_Union";  | 
|
750  | 
val Transset_Union_family = thm "Transset_Union_family";  | 
|
751  | 
val Transset_Inter_family = thm "Transset_Inter_family";  | 
|
752  | 
val OrdI = thm "OrdI";  | 
|
753  | 
val Ord_is_Transset = thm "Ord_is_Transset";  | 
|
754  | 
val Ord_contains_Transset = thm "Ord_contains_Transset";  | 
|
755  | 
val Ord_in_Ord = thm "Ord_in_Ord";  | 
|
756  | 
val Ord_succD = thm "Ord_succD";  | 
|
757  | 
val Ord_subset_Ord = thm "Ord_subset_Ord";  | 
|
758  | 
val OrdmemD = thm "OrdmemD";  | 
|
759  | 
val Ord_trans = thm "Ord_trans";  | 
|
760  | 
val Ord_succ_subsetI = thm "Ord_succ_subsetI";  | 
|
761  | 
val Ord_0 = thm "Ord_0";  | 
|
762  | 
val Ord_succ = thm "Ord_succ";  | 
|
763  | 
val Ord_1 = thm "Ord_1";  | 
|
764  | 
val Ord_succ_iff = thm "Ord_succ_iff";  | 
|
765  | 
val Ord_Un = thm "Ord_Un";  | 
|
766  | 
val Ord_Int = thm "Ord_Int";  | 
|
767  | 
val Ord_Inter = thm "Ord_Inter";  | 
|
768  | 
val Ord_INT = thm "Ord_INT";  | 
|
769  | 
val ON_class = thm "ON_class";  | 
|
770  | 
val ltI = thm "ltI";  | 
|
771  | 
val ltE = thm "ltE";  | 
|
772  | 
val ltD = thm "ltD";  | 
|
773  | 
val not_lt0 = thm "not_lt0";  | 
|
774  | 
val lt_Ord = thm "lt_Ord";  | 
|
775  | 
val lt_Ord2 = thm "lt_Ord2";  | 
|
776  | 
val le_Ord2 = thm "le_Ord2";  | 
|
777  | 
val lt0E = thm "lt0E";  | 
|
778  | 
val lt_trans = thm "lt_trans";  | 
|
779  | 
val lt_not_sym = thm "lt_not_sym";  | 
|
780  | 
val lt_asym = thm "lt_asym";  | 
|
781  | 
val lt_irrefl = thm "lt_irrefl";  | 
|
782  | 
val lt_not_refl = thm "lt_not_refl";  | 
|
783  | 
val le_iff = thm "le_iff";  | 
|
784  | 
val leI = thm "leI";  | 
|
785  | 
val le_eqI = thm "le_eqI";  | 
|
786  | 
val le_refl = thm "le_refl";  | 
|
787  | 
val le_refl_iff = thm "le_refl_iff";  | 
|
788  | 
val leCI = thm "leCI";  | 
|
789  | 
val leE = thm "leE";  | 
|
790  | 
val le_anti_sym = thm "le_anti_sym";  | 
|
791  | 
val le0_iff = thm "le0_iff";  | 
|
792  | 
val le0D = thm "le0D";  | 
|
793  | 
val Memrel_iff = thm "Memrel_iff";  | 
|
794  | 
val MemrelI = thm "MemrelI";  | 
|
795  | 
val MemrelE = thm "MemrelE";  | 
|
796  | 
val Memrel_type = thm "Memrel_type";  | 
|
797  | 
val Memrel_mono = thm "Memrel_mono";  | 
|
798  | 
val Memrel_0 = thm "Memrel_0";  | 
|
799  | 
val Memrel_1 = thm "Memrel_1";  | 
|
800  | 
val wf_Memrel = thm "wf_Memrel";  | 
|
801  | 
val trans_Memrel = thm "trans_Memrel";  | 
|
802  | 
val Transset_Memrel_iff = thm "Transset_Memrel_iff";  | 
|
803  | 
val Transset_induct = thm "Transset_induct";  | 
|
804  | 
val Ord_induct = thm "Ord_induct";  | 
|
805  | 
val trans_induct = thm "trans_induct";  | 
|
806  | 
val Ord_linear = thm "Ord_linear";  | 
|
807  | 
val Ord_linear_lt = thm "Ord_linear_lt";  | 
|
808  | 
val Ord_linear2 = thm "Ord_linear2";  | 
|
809  | 
val Ord_linear_le = thm "Ord_linear_le";  | 
|
810  | 
val le_imp_not_lt = thm "le_imp_not_lt";  | 
|
811  | 
val not_lt_imp_le = thm "not_lt_imp_le";  | 
|
812  | 
val Ord_mem_iff_lt = thm "Ord_mem_iff_lt";  | 
|
813  | 
val not_lt_iff_le = thm "not_lt_iff_le";  | 
|
814  | 
val not_le_iff_lt = thm "not_le_iff_lt";  | 
|
815  | 
val Ord_0_le = thm "Ord_0_le";  | 
|
816  | 
val Ord_0_lt = thm "Ord_0_lt";  | 
|
817  | 
val Ord_0_lt_iff = thm "Ord_0_lt_iff";  | 
|
818  | 
val zero_le_succ_iff = thm "zero_le_succ_iff";  | 
|
819  | 
val subset_imp_le = thm "subset_imp_le";  | 
|
820  | 
val le_imp_subset = thm "le_imp_subset";  | 
|
821  | 
val le_subset_iff = thm "le_subset_iff";  | 
|
822  | 
val le_succ_iff = thm "le_succ_iff";  | 
|
823  | 
val all_lt_imp_le = thm "all_lt_imp_le";  | 
|
824  | 
val lt_trans1 = thm "lt_trans1";  | 
|
825  | 
val lt_trans2 = thm "lt_trans2";  | 
|
826  | 
val le_trans = thm "le_trans";  | 
|
827  | 
val succ_leI = thm "succ_leI";  | 
|
828  | 
val succ_leE = thm "succ_leE";  | 
|
829  | 
val succ_le_iff = thm "succ_le_iff";  | 
|
830  | 
val succ_le_imp_le = thm "succ_le_imp_le";  | 
|
831  | 
val lt_subset_trans = thm "lt_subset_trans";  | 
|
832  | 
val Un_upper1_le = thm "Un_upper1_le";  | 
|
833  | 
val Un_upper2_le = thm "Un_upper2_le";  | 
|
834  | 
val Un_least_lt = thm "Un_least_lt";  | 
|
835  | 
val Un_least_lt_iff = thm "Un_least_lt_iff";  | 
|
836  | 
val Un_least_mem_iff = thm "Un_least_mem_iff";  | 
|
837  | 
val Int_greatest_lt = thm "Int_greatest_lt";  | 
|
838  | 
val Ord_Union = thm "Ord_Union";  | 
|
839  | 
val Ord_UN = thm "Ord_UN";  | 
|
840  | 
val UN_least_le = thm "UN_least_le";  | 
|
841  | 
val UN_succ_least_lt = thm "UN_succ_least_lt";  | 
|
842  | 
val UN_upper_le = thm "UN_upper_le";  | 
|
843  | 
val le_implies_UN_le_UN = thm "le_implies_UN_le_UN";  | 
|
844  | 
val Ord_equality = thm "Ord_equality";  | 
|
845  | 
val Ord_Union_subset = thm "Ord_Union_subset";  | 
|
846  | 
val Limit_Union_eq = thm "Limit_Union_eq";  | 
|
847  | 
val Limit_is_Ord = thm "Limit_is_Ord";  | 
|
848  | 
val Limit_has_0 = thm "Limit_has_0";  | 
|
849  | 
val Limit_has_succ = thm "Limit_has_succ";  | 
|
850  | 
val non_succ_LimitI = thm "non_succ_LimitI";  | 
|
851  | 
val succ_LimitE = thm "succ_LimitE";  | 
|
852  | 
val not_succ_Limit = thm "not_succ_Limit";  | 
|
853  | 
val Limit_le_succD = thm "Limit_le_succD";  | 
|
854  | 
val Ord_cases_disj = thm "Ord_cases_disj";  | 
|
855  | 
val Ord_cases = thm "Ord_cases";  | 
|
856  | 
val trans_induct3 = thm "trans_induct3";  | 
|
857  | 
*}  | 
|
| 435 | 858  | 
|
859  | 
end  |