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