author | wenzelm |
Sun, 10 Nov 2024 12:23:41 +0100 | |
changeset 81418 | de8dbdadda76 |
parent 81181 | ff2a637a449e |
permissions | -rw-r--r-- |
13505 | 1 |
(* Title: ZF/Constructible/Normal.thy |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
3 |
*) |
|
4 |
||
60770 | 5 |
section \<open>Closed Unbounded Classes and Normal Functions\<close> |
13223 | 6 |
|
65449
c82e63b11b8b
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents:
61798
diff
changeset
|
7 |
theory Normal imports ZF begin |
13223 | 8 |
|
60770 | 9 |
text\<open> |
13223 | 10 |
One source is the book |
11 |
||
12 |
Frank R. Drake. |
|
13 |
\emph{Set Theory: An Introduction to Large Cardinals}. |
|
14 |
North-Holland, 1974. |
|
60770 | 15 |
\<close> |
13223 | 16 |
|
17 |
||
60770 | 18 |
subsection \<open>Closed and Unbounded (c.u.) Classes of Ordinals\<close> |
13223 | 19 |
|
21233 | 20 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
21 |
Closed :: "(i\<Rightarrow>o) \<Rightarrow> o" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
22 |
"Closed(P) \<equiv> \<forall>I. I \<noteq> 0 \<longrightarrow> (\<forall>i\<in>I. Ord(i) \<and> P(i)) \<longrightarrow> P(\<Union>(I))" |
13223 | 23 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
24 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
25 |
Unbounded :: "(i\<Rightarrow>o) \<Rightarrow> o" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
26 |
"Unbounded(P) \<equiv> \<forall>i. Ord(i) \<longrightarrow> (\<exists>j. i<j \<and> P(j))" |
13223 | 27 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
28 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
29 |
Closed_Unbounded :: "(i\<Rightarrow>o) \<Rightarrow> o" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
30 |
"Closed_Unbounded(P) \<equiv> Closed(P) \<and> Unbounded(P)" |
13223 | 31 |
|
32 |
||
60770 | 33 |
subsubsection\<open>Simple facts about c.u. classes\<close> |
13223 | 34 |
|
35 |
lemma ClosedI: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
36 |
"\<lbrakk>\<And>I. \<lbrakk>I \<noteq> 0; \<forall>i\<in>I. Ord(i) \<and> P(i)\<rbrakk> \<Longrightarrow> P(\<Union>(I))\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
37 |
\<Longrightarrow> Closed(P)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
38 |
by (simp add: Closed_def) |
13223 | 39 |
|
40 |
lemma ClosedD: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
41 |
"\<lbrakk>Closed(P); I \<noteq> 0; \<And>i. i\<in>I \<Longrightarrow> Ord(i); \<And>i. i\<in>I \<Longrightarrow> P(i)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
42 |
\<Longrightarrow> P(\<Union>(I))" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
43 |
by (simp add: Closed_def) |
13223 | 44 |
|
45 |
lemma UnboundedD: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
46 |
"\<lbrakk>Unbounded(P); Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. i<j \<and> P(j)" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
47 |
by (simp add: Unbounded_def) |
13223 | 48 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
49 |
lemma Closed_Unbounded_imp_Unbounded: "Closed_Unbounded(C) \<Longrightarrow> Unbounded(C)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
50 |
by (simp add: Closed_Unbounded_def) |
13223 | 51 |
|
52 |
||
60770 | 53 |
text\<open>The universal class, V, is closed and unbounded. |
54 |
A bit odd, since C. U. concerns only ordinals, but it's used below!\<close> |
|
13223 | 55 |
theorem Closed_Unbounded_V [simp]: "Closed_Unbounded(\<lambda>x. True)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
56 |
by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast) |
13223 | 57 |
|
69593 | 58 |
text\<open>The class of ordinals, \<^term>\<open>Ord\<close>, is closed and unbounded.\<close> |
13223 | 59 |
theorem Closed_Unbounded_Ord [simp]: "Closed_Unbounded(Ord)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
60 |
by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast) |
13223 | 61 |
|
69593 | 62 |
text\<open>The class of limit ordinals, \<^term>\<open>Limit\<close>, is closed and unbounded.\<close> |
13223 | 63 |
theorem Closed_Unbounded_Limit [simp]: "Closed_Unbounded(Limit)" |
76218
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
64 |
proof - |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
65 |
have "\<exists>j. i < j \<and> Limit(j)" if "Ord(i)" for i |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
66 |
apply (rule_tac x="i++nat" in exI) |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
67 |
apply (blast intro: oadd_lt_self oadd_LimitI Limit_has_0 that) |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
68 |
done |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
69 |
then show ?thesis |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
70 |
by (auto simp: Closed_Unbounded_def Closed_def Unbounded_def Limit_Union) |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
71 |
qed |
13223 | 72 |
|
69593 | 73 |
text\<open>The class of cardinals, \<^term>\<open>Card\<close>, is closed and unbounded.\<close> |
13223 | 74 |
theorem Closed_Unbounded_Card [simp]: "Closed_Unbounded(Card)" |
76218
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
75 |
proof - |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
76 |
have "\<forall>i. Ord(i) \<longrightarrow> (\<exists>j. i < j \<and> Card(j))" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
77 |
by (blast intro: lt_csucc Card_csucc) |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
78 |
then show ?thesis |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
79 |
by (simp add: Closed_Unbounded_def Closed_def Unbounded_def) |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
80 |
qed |
13223 | 81 |
|
82 |
||
60770 | 83 |
subsubsection\<open>The intersection of any set-indexed family of c.u. classes is |
84 |
c.u.\<close> |
|
13223 | 85 |
|
60770 | 86 |
text\<open>The constructions below come from Kunen, \emph{Set Theory}, page 78.\<close> |
13428 | 87 |
locale cub_family = |
13223 | 88 |
fixes P and A |
69593 | 89 |
fixes next_greater \<comment> \<open>the next ordinal satisfying class \<^term>\<open>A\<close>\<close> |
90 |
fixes sup_greater \<comment> \<open>sup of those ordinals over all \<^term>\<open>A\<close>\<close> |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
91 |
assumes closed: "a\<in>A \<Longrightarrow> Closed(P(a))" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
92 |
and unbounded: "a\<in>A \<Longrightarrow> Unbounded(P(a))" |
13223 | 93 |
and A_non0: "A\<noteq>0" |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
94 |
defines "next_greater(a,x) \<equiv> \<mu> y. x<y \<and> P(a,y)" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
95 |
and "sup_greater(x) \<equiv> \<Union>a\<in>A. next_greater(a,x)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
96 |
|
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
97 |
begin |
13223 | 98 |
|
60770 | 99 |
text\<open>Trivial that the intersection is closed.\<close> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
100 |
lemma Closed_INT: "Closed(\<lambda>x. \<forall>i\<in>A. P(i,x))" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
101 |
by (blast intro: ClosedI ClosedD [OF closed]) |
13223 | 102 |
|
60770 | 103 |
text\<open>All remaining effort goes to show that the intersection is unbounded.\<close> |
13223 | 104 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
105 |
lemma Ord_sup_greater: |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
106 |
"Ord(sup_greater(x))" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
107 |
by (simp add: sup_greater_def next_greater_def) |
13223 | 108 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
109 |
lemma Ord_next_greater: |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
110 |
"Ord(next_greater(a,x))" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
111 |
by (simp add: next_greater_def) |
13223 | 112 |
|
69593 | 113 |
text\<open>\<^term>\<open>next_greater\<close> works as expected: it returns a larger value |
114 |
and one that belongs to class \<^term>\<open>P(a)\<close>.\<close> |
|
76218
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
115 |
lemma |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
116 |
assumes "Ord(x)" "a\<in>A" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
117 |
shows next_greater_in_P: "P(a, next_greater(a,x))" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
118 |
and next_greater_gt: "x < next_greater(a,x)" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
119 |
proof - |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
120 |
obtain y where "x < y" "P(a,y)" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
121 |
using assms UnboundedD [OF unbounded] by blast |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
122 |
then have "P(a, next_greater(a,x)) \<and> x < next_greater(a,x)" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
123 |
unfolding next_greater_def |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
124 |
by (blast intro: LeastI2 lt_Ord2) |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
125 |
then show "P(a, next_greater(a,x))" "x < next_greater(a,x)" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
126 |
by auto |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
127 |
qed |
13223 | 128 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
129 |
lemma sup_greater_gt: |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
130 |
"Ord(x) \<Longrightarrow> x < sup_greater(x)" |
76218
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
131 |
using A_non0 unfolding sup_greater_def |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
132 |
by (blast intro: UN_upper_lt next_greater_gt Ord_next_greater) |
13223 | 133 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
134 |
lemma next_greater_le_sup_greater: |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
135 |
"a\<in>A \<Longrightarrow> next_greater(a,x) \<le> sup_greater(x)" |
76218
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
136 |
unfolding sup_greater_def |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
137 |
by (force intro: UN_upper_le Ord_next_greater) |
13223 | 138 |
|
76218
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
139 |
lemma omega_sup_greater_eq_UN: |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
140 |
assumes "Ord(x)" "a\<in>A" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
141 |
shows "sup_greater^\<omega> (x) = |
13223 | 142 |
(\<Union>n\<in>nat. next_greater(a, sup_greater^n (x)))" |
76218
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
143 |
proof (rule le_anti_sym) |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
144 |
show "sup_greater^\<omega> (x) \<le> (\<Union>n\<in>nat. next_greater(a, sup_greater^n (x)))" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
145 |
using assms |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
146 |
unfolding iterates_omega_def |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
147 |
by (blast intro: leI le_implies_UN_le_UN next_greater_gt Ord_iterates Ord_sup_greater) |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
148 |
next |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
149 |
have "Ord(\<Union>n\<in>nat. sup_greater^n (x))" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
150 |
by (blast intro: Ord_iterates Ord_sup_greater assms) |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
151 |
moreover have "next_greater(a, sup_greater^n (x)) \<le> |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
152 |
(\<Union>n\<in>nat. sup_greater^n (x))" if "n \<in> nat" for n |
76221 | 153 |
proof (rule UN_upper_le) |
154 |
show "next_greater(a, sup_greater^n (x)) \<le> sup_greater^succ(n) (x)" |
|
155 |
using assms by (simp add: next_greater_le_sup_greater) |
|
156 |
show "Ord(\<Union>xa\<in>nat. sup_greater^xa (x))" |
|
157 |
using assms by (blast intro: Ord_iterates Ord_sup_greater) |
|
158 |
qed (use that in auto) |
|
76218
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
159 |
ultimately |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
160 |
show "(\<Union>n\<in>nat. next_greater(a, sup_greater^n (x))) \<le> sup_greater^\<omega> (x)" |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
161 |
using assms unfolding iterates_omega_def by (blast intro: UN_least_le) |
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
162 |
qed |
13223 | 163 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
164 |
lemma P_omega_sup_greater: |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
165 |
"\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> P(a, sup_greater^\<omega> (x))" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
166 |
apply (simp add: omega_sup_greater_eq_UN) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
167 |
apply (rule ClosedD [OF closed]) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
168 |
apply (blast intro: ltD, auto) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
169 |
apply (blast intro: Ord_iterates Ord_next_greater Ord_sup_greater) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
170 |
apply (blast intro: next_greater_in_P Ord_iterates Ord_sup_greater) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
171 |
done |
13223 | 172 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
173 |
lemma omega_sup_greater_gt: |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
174 |
"Ord(x) \<Longrightarrow> x < sup_greater^\<omega> (x)" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
175 |
apply (simp add: iterates_omega_def) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
176 |
apply (rule UN_upper_lt [of 1], simp_all) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
177 |
apply (blast intro: sup_greater_gt) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
178 |
apply (blast intro: Ord_iterates Ord_sup_greater) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
179 |
done |
13223 | 180 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
181 |
lemma Unbounded_INT: "Unbounded(\<lambda>x. \<forall>a\<in>A. P(a,x))" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
182 |
unfolding Unbounded_def |
76218
728f38b016c0
added a couple of structured proofs
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
183 |
by (blast intro!: omega_sup_greater_gt P_omega_sup_greater) |
13223 | 184 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
185 |
lemma Closed_Unbounded_INT: |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
186 |
"Closed_Unbounded(\<lambda>x. \<forall>a\<in>A. P(a,x))" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
187 |
by (simp add: Closed_Unbounded_def Closed_INT Unbounded_INT) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
188 |
|
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
189 |
end |
13223 | 190 |
|
191 |
||
192 |
theorem Closed_Unbounded_INT: |
|
76221 | 193 |
assumes "\<And>a. a\<in>A \<Longrightarrow> Closed_Unbounded(P(a))" |
194 |
shows "Closed_Unbounded(\<lambda>x. \<forall>a\<in>A. P(a, x))" |
|
195 |
proof (cases "A=0") |
|
196 |
case False |
|
197 |
with assms [unfolded Closed_Unbounded_def] show ?thesis |
|
198 |
by (intro cub_family.Closed_Unbounded_INT [OF cub_family.intro]) auto |
|
199 |
qed auto |
|
13223 | 200 |
|
201 |
lemma Int_iff_INT2: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
202 |
"P(x) \<and> Q(x) \<longleftrightarrow> (\<forall>i\<in>2. (i=0 \<longrightarrow> P(x)) \<and> (i=1 \<longrightarrow> Q(x)))" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
203 |
by auto |
13223 | 204 |
|
205 |
theorem Closed_Unbounded_Int: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
206 |
"\<lbrakk>Closed_Unbounded(P); Closed_Unbounded(Q)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
207 |
\<Longrightarrow> Closed_Unbounded(\<lambda>x. P(x) \<and> Q(x))" |
76221 | 208 |
unfolding Int_iff_INT2 |
209 |
by (rule Closed_Unbounded_INT, auto) |
|
210 |
||
13223 | 211 |
|
212 |
||
60770 | 213 |
subsection \<open>Normal Functions\<close> |
13223 | 214 |
|
21233 | 215 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
216 |
mono_le_subset :: "(i\<Rightarrow>i) \<Rightarrow> o" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
217 |
"mono_le_subset(M) \<equiv> \<forall>i j. i\<le>j \<longrightarrow> M(i) \<subseteq> M(j)" |
13223 | 218 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
219 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
220 |
mono_Ord :: "(i\<Rightarrow>i) \<Rightarrow> o" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
221 |
"mono_Ord(F) \<equiv> \<forall>i j. i<j \<longrightarrow> F(i) < F(j)" |
13223 | 222 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
223 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
224 |
cont_Ord :: "(i\<Rightarrow>i) \<Rightarrow> o" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
225 |
"cont_Ord(F) \<equiv> \<forall>l. Limit(l) \<longrightarrow> F(l) = (\<Union>i<l. F(i))" |
13223 | 226 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
227 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
228 |
Normal :: "(i\<Rightarrow>i) \<Rightarrow> o" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
229 |
"Normal(F) \<equiv> mono_Ord(F) \<and> cont_Ord(F)" |
13223 | 230 |
|
231 |
||
60770 | 232 |
subsubsection\<open>Immediate properties of the definitions\<close> |
13223 | 233 |
|
234 |
lemma NormalI: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
235 |
"\<lbrakk>\<And>i j. i<j \<Longrightarrow> F(i) < F(j); \<And>l. Limit(l) \<Longrightarrow> F(l) = (\<Union>i<l. F(i))\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
236 |
\<Longrightarrow> Normal(F)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
237 |
by (simp add: Normal_def mono_Ord_def cont_Ord_def) |
13223 | 238 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
239 |
lemma mono_Ord_imp_Ord: "\<lbrakk>Ord(i); mono_Ord(F)\<rbrakk> \<Longrightarrow> Ord(F(i))" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
240 |
apply (auto simp add: mono_Ord_def) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
241 |
apply (blast intro: lt_Ord) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
242 |
done |
13223 | 243 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
244 |
lemma mono_Ord_imp_mono: "\<lbrakk>i<j; mono_Ord(F)\<rbrakk> \<Longrightarrow> F(i) < F(j)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
245 |
by (simp add: mono_Ord_def) |
13223 | 246 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
247 |
lemma Normal_imp_Ord [simp]: "\<lbrakk>Normal(F); Ord(i)\<rbrakk> \<Longrightarrow> Ord(F(i))" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
248 |
by (simp add: Normal_def mono_Ord_imp_Ord) |
13223 | 249 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
250 |
lemma Normal_imp_cont: "\<lbrakk>Normal(F); Limit(l)\<rbrakk> \<Longrightarrow> F(l) = (\<Union>i<l. F(i))" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
251 |
by (simp add: Normal_def cont_Ord_def) |
13223 | 252 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
253 |
lemma Normal_imp_mono: "\<lbrakk>i<j; Normal(F)\<rbrakk> \<Longrightarrow> F(i) < F(j)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
254 |
by (simp add: Normal_def mono_Ord_def) |
13223 | 255 |
|
46963 | 256 |
lemma Normal_increasing: |
257 |
assumes i: "Ord(i)" and F: "Normal(F)" shows"i \<le> F(i)" |
|
258 |
using i |
|
259 |
proof (induct i rule: trans_induct3) |
|
260 |
case 0 thus ?case by (simp add: subset_imp_le F) |
|
261 |
next |
|
262 |
case (succ i) |
|
263 |
hence "F(i) < F(succ(i))" using F |
|
264 |
by (simp add: Normal_def mono_Ord_def) |
|
265 |
thus ?case using succ.hyps |
|
266 |
by (blast intro: lt_trans1) |
|
267 |
next |
|
268 |
case (limit l) |
|
269 |
hence "l = (\<Union>y<l. y)" |
|
270 |
by (simp add: Limit_OUN_eq) |
|
271 |
also have "... \<le> (\<Union>y<l. F(y))" using limit |
|
272 |
by (blast intro: ltD le_implies_OUN_le_OUN) |
|
273 |
finally have "l \<le> (\<Union>y<l. F(y))" . |
|
274 |
moreover have "(\<Union>y<l. F(y)) \<le> F(l)" using limit F |
|
275 |
by (simp add: Normal_imp_cont lt_Ord) |
|
276 |
ultimately show ?case |
|
277 |
by (blast intro: le_trans) |
|
278 |
qed |
|
13223 | 279 |
|
280 |
||
60770 | 281 |
subsubsection\<open>The class of fixedpoints is closed and unbounded\<close> |
13223 | 282 |
|
60770 | 283 |
text\<open>The proof is from Drake, pages 113--114.\<close> |
13223 | 284 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
285 |
lemma mono_Ord_imp_le_subset: "mono_Ord(F) \<Longrightarrow> mono_le_subset(F)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
286 |
apply (simp add: mono_le_subset_def, clarify) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
287 |
apply (subgoal_tac "F(i)\<le>F(j)", blast dest: le_imp_subset) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
288 |
apply (simp add: le_iff) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
289 |
apply (blast intro: lt_Ord2 mono_Ord_imp_Ord mono_Ord_imp_mono) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
290 |
done |
13223 | 291 |
|
60770 | 292 |
text\<open>The following equation is taken for granted in any set theory text.\<close> |
13223 | 293 |
lemma cont_Ord_Union: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
294 |
"\<lbrakk>cont_Ord(F); mono_le_subset(F); X=0 \<longrightarrow> F(0)=0; \<forall>x\<in>X. Ord(x)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
295 |
\<Longrightarrow> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
296 |
apply (frule Ord_set_cases) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
297 |
apply (erule disjE, force) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
298 |
apply (thin_tac "X=0 \<longrightarrow> Q" for Q, auto) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
299 |
txt\<open>The trival case of \<^term>\<open>\<Union>X \<in> X\<close>\<close> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
300 |
apply (rule equalityI, blast intro: Ord_Union_eq_succD) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
301 |
apply (simp add: mono_le_subset_def UN_subset_iff le_subset_iff) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
302 |
apply (blast elim: equalityE) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
303 |
txt\<open>The limit case, \<^term>\<open>Limit(\<Union>X)\<close>: |
13223 | 304 |
@{subgoals[display,indent=0,margin=65]} |
60770 | 305 |
\<close> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
306 |
apply (simp add: OUN_Union_eq cont_Ord_def) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
307 |
apply (rule equalityI) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
308 |
txt\<open>First inclusion:\<close> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
309 |
apply (rule UN_least [OF OUN_least]) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
310 |
apply (simp add: mono_le_subset_def, blast intro: leI) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
311 |
txt\<open>Second inclusion:\<close> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
312 |
apply (rule UN_least) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
313 |
apply (frule Union_upper_le, blast, blast) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
314 |
apply (erule leE, drule ltD, elim UnionE) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
315 |
apply (simp add: OUnion_def) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
316 |
apply blast+ |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
317 |
done |
13223 | 318 |
|
319 |
lemma Normal_Union: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
320 |
"\<lbrakk>X\<noteq>0; \<forall>x\<in>X. Ord(x); Normal(F)\<rbrakk> \<Longrightarrow> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))" |
76221 | 321 |
unfolding Normal_def |
322 |
by (blast intro: mono_Ord_imp_le_subset cont_Ord_Union) |
|
323 |
||
13223 | 324 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
325 |
lemma Normal_imp_fp_Closed: "Normal(F) \<Longrightarrow> Closed(\<lambda>i. F(i) = i)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
326 |
apply (simp add: Closed_def ball_conj_distrib, clarify) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
327 |
apply (frule Ord_set_cases) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
328 |
apply (auto simp add: Normal_Union) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
329 |
done |
13223 | 330 |
|
331 |
||
332 |
lemma iterates_Normal_increasing: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
333 |
"\<lbrakk>n\<in>nat; x < F(x); Normal(F)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
334 |
\<Longrightarrow> F^n (x) < F^(succ(n)) (x)" |
76221 | 335 |
by (induct n rule: nat_induct) (simp_all add: Normal_imp_mono) |
13223 | 336 |
|
337 |
lemma Ord_iterates_Normal: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
338 |
"\<lbrakk>n\<in>nat; Normal(F); Ord(x)\<rbrakk> \<Longrightarrow> Ord(F^n (x))" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
339 |
by (simp) |
13223 | 340 |
|
60770 | 341 |
text\<open>THIS RESULT IS UNUSED\<close> |
13223 | 342 |
lemma iterates_omega_Limit: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
343 |
"\<lbrakk>Normal(F); x < F(x)\<rbrakk> \<Longrightarrow> Limit(F^\<omega> (x))" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
344 |
apply (frule lt_Ord) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
345 |
apply (simp add: iterates_omega_def) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
346 |
apply (rule increasing_LimitI) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
347 |
\<comment> \<open>this lemma is @{thm increasing_LimitI [no_vars]}\<close> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
348 |
apply (blast intro: UN_upper_lt [of "1"] Normal_imp_Ord |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
349 |
Ord_iterates lt_imp_0_lt |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
350 |
iterates_Normal_increasing, clarify) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
351 |
apply (rule bexI) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
352 |
apply (blast intro: Ord_in_Ord [OF Ord_iterates_Normal]) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
353 |
apply (rule UN_I, erule nat_succI) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
354 |
apply (blast intro: iterates_Normal_increasing Ord_iterates_Normal |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
355 |
ltD [OF lt_trans1, OF succ_leI, OF ltI]) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
356 |
done |
13223 | 357 |
|
358 |
lemma iterates_omega_fixedpoint: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
359 |
"\<lbrakk>Normal(F); Ord(a)\<rbrakk> \<Longrightarrow> F(F^\<omega> (a)) = F^\<omega> (a)" |
13223 | 360 |
apply (frule Normal_increasing, assumption) |
361 |
apply (erule leE) |
|
362 |
apply (simp_all add: iterates_omega_triv [OF sym]) (*for subgoal 2*) |
|
363 |
apply (simp add: iterates_omega_def Normal_Union) |
|
364 |
apply (rule equalityI, force simp add: nat_succI) |
|
60770 | 365 |
txt\<open>Opposite inclusion: |
13223 | 366 |
@{subgoals[display,indent=0,margin=65]} |
60770 | 367 |
\<close> |
13223 | 368 |
apply clarify |
369 |
apply (rule UN_I, assumption) |
|
370 |
apply (frule iterates_Normal_increasing, assumption, assumption, simp) |
|
371 |
apply (blast intro: Ord_trans ltD Ord_iterates_Normal Normal_imp_Ord [of F]) |
|
372 |
done |
|
373 |
||
374 |
lemma iterates_omega_increasing: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
375 |
"\<lbrakk>Normal(F); Ord(a)\<rbrakk> \<Longrightarrow> a \<le> F^\<omega> (a)" |
76221 | 376 |
by (simp add: iterates_omega_def UN_upper_le [of 0]) |
13223 | 377 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
378 |
lemma Normal_imp_fp_Unbounded: "Normal(F) \<Longrightarrow> Unbounded(\<lambda>i. F(i) = i)" |
13223 | 379 |
apply (unfold Unbounded_def, clarify) |
380 |
apply (rule_tac x="F^\<omega> (succ(i))" in exI) |
|
381 |
apply (simp add: iterates_omega_fixedpoint) |
|
382 |
apply (blast intro: lt_trans2 [OF _ iterates_omega_increasing]) |
|
383 |
done |
|
384 |
||
385 |
||
386 |
theorem Normal_imp_fp_Closed_Unbounded: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
387 |
"Normal(F) \<Longrightarrow> Closed_Unbounded(\<lambda>i. F(i) = i)" |
76221 | 388 |
by (simp add: Closed_Unbounded_def Normal_imp_fp_Closed Normal_imp_fp_Unbounded) |
13223 | 389 |
|
390 |
||
61798 | 391 |
subsubsection\<open>Function \<open>normalize\<close>\<close> |
13223 | 392 |
|
61798 | 393 |
text\<open>Function \<open>normalize\<close> maps a function \<open>F\<close> to a |
13223 | 394 |
normal function that bounds it above. The result is normal if and |
61798 | 395 |
only if \<open>F\<close> is continuous: succ is not bounded above by any |
13223 | 396 |
normal function, by @{thm [source] Normal_imp_fp_Unbounded}. |
60770 | 397 |
\<close> |
21233 | 398 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
399 |
normalize :: "[i\<Rightarrow>i, i] \<Rightarrow> i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
400 |
"normalize(F,a) \<equiv> transrec2(a, F(0), \<lambda>x r. F(succ(x)) \<union> succ(r))" |
13223 | 401 |
|
402 |
||
403 |
lemma Ord_normalize [simp, intro]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
404 |
"\<lbrakk>Ord(a); \<And>x. Ord(x) \<Longrightarrow> Ord(F(x))\<rbrakk> \<Longrightarrow> Ord(normalize(F, a))" |
76221 | 405 |
proof (induct a rule: trans_induct3) |
406 |
qed (simp_all add: ltD def_transrec2 [OF normalize_def]) |
|
13223 | 407 |
|
408 |
lemma normalize_increasing: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
409 |
assumes ab: "a < b" and F: "\<And>x. Ord(x) \<Longrightarrow> Ord(F(x))" |
46963 | 410 |
shows "normalize(F,a) < normalize(F,b)" |
411 |
proof - |
|
81181 | 412 |
have "Ord(b)" using ab by (blast intro: lt_Ord2) |
413 |
hence "x < b \<Longrightarrow> normalize(F,x) < normalize(F,b)" for x |
|
414 |
proof (induct b arbitrary: x rule: trans_induct3) |
|
415 |
case 0 thus ?case by simp |
|
416 |
next |
|
417 |
case (succ b) |
|
418 |
thus ?case |
|
419 |
by (auto simp add: le_iff def_transrec2 [OF normalize_def] intro: Un_upper2_lt F) |
|
420 |
next |
|
421 |
case (limit l) |
|
422 |
hence sc: "succ(x) < l" |
|
423 |
by (blast intro: Limit_has_succ) |
|
424 |
hence "normalize(F,x) < normalize(F,succ(x))" |
|
425 |
by (blast intro: limit elim: ltE) |
|
426 |
hence "normalize(F,x) < (\<Union>j<l. normalize(F,j))" |
|
427 |
by (blast intro: OUN_upper_lt lt_Ord F sc) |
|
428 |
thus ?case using limit |
|
429 |
by (simp add: def_transrec2 [OF normalize_def]) |
|
430 |
qed |
|
431 |
thus ?thesis using ab . |
|
46963 | 432 |
qed |
13223 | 433 |
|
434 |
theorem Normal_normalize: |
|
76221 | 435 |
assumes "\<And>x. Ord(x) \<Longrightarrow> Ord(F(x))" shows "Normal(normalize(F))" |
436 |
proof (rule NormalI) |
|
437 |
show "\<And>i j. i < j \<Longrightarrow> normalize(F,i) < normalize(F,j)" |
|
438 |
using assms by (blast intro!: normalize_increasing) |
|
439 |
show "\<And>l. Limit(l) \<Longrightarrow> normalize(F, l) = (\<Union>i<l. normalize(F,i))" |
|
440 |
by (simp add: def_transrec2 [OF normalize_def]) |
|
441 |
qed |
|
13223 | 442 |
|
443 |
theorem le_normalize: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
444 |
assumes a: "Ord(a)" and coF: "cont_Ord(F)" and F: "\<And>x. Ord(x) \<Longrightarrow> Ord(F(x))" |
46963 | 445 |
shows "F(a) \<le> normalize(F,a)" |
446 |
using a |
|
447 |
proof (induct a rule: trans_induct3) |
|
448 |
case 0 thus ?case by (simp add: F def_transrec2 [OF normalize_def]) |
|
449 |
next |
|
450 |
case (succ a) |
|
451 |
thus ?case |
|
452 |
by (simp add: def_transrec2 [OF normalize_def] Un_upper1_le F ) |
|
453 |
next |
|
454 |
case (limit l) |
|
455 |
thus ?case using F coF [unfolded cont_Ord_def] |
|
456 |
by (simp add: def_transrec2 [OF normalize_def] le_implies_OUN_le_OUN ltD) |
|
457 |
qed |
|
13223 | 458 |
|
459 |
||
60770 | 460 |
subsection \<open>The Alephs\<close> |
461 |
text \<open>This is the well-known transfinite enumeration of the cardinal |
|
462 |
numbers.\<close> |
|
13223 | 463 |
|
21233 | 464 |
definition |
81125 | 465 |
Aleph :: "i \<Rightarrow> i" (\<open>(\<open>open_block notation=\<open>prefix \<aleph>\<close>\<close>\<aleph>_)\<close> [90] 90) where |
81181 | 466 |
"\<aleph>a \<equiv> transrec2(a, nat, \<lambda>x r. csucc(r))" |
13223 | 467 |
|
468 |
lemma Card_Aleph [simp, intro]: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
469 |
"Ord(a) \<Longrightarrow> Card(Aleph(a))" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
470 |
apply (erule trans_induct3) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
471 |
apply (simp_all add: Card_csucc Card_nat Card_is_Ord |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
472 |
def_transrec2 [OF Aleph_def]) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
473 |
done |
13223 | 474 |
|
475 |
lemma Aleph_increasing: |
|
46963 | 476 |
assumes ab: "a < b" shows "Aleph(a) < Aleph(b)" |
477 |
proof - |
|
81181 | 478 |
have "Ord(b)" using ab by (blast intro: lt_Ord2) |
479 |
hence "x < b \<Longrightarrow> Aleph(x) < Aleph(b)" for x |
|
480 |
proof (induct b arbitrary: x rule: trans_induct3) |
|
481 |
case 0 thus ?case by simp |
|
482 |
next |
|
483 |
case (succ b) |
|
484 |
thus ?case |
|
485 |
by (force simp add: le_iff def_transrec2 [OF Aleph_def] |
|
486 |
intro: lt_trans lt_csucc Card_is_Ord) |
|
487 |
next |
|
488 |
case (limit l) |
|
489 |
hence sc: "succ(x) < l" |
|
490 |
by (blast intro: Limit_has_succ) |
|
491 |
hence "\<aleph>x < (\<Union>j<l. \<aleph>j)" using limit |
|
492 |
by (blast intro: OUN_upper_lt Card_is_Ord ltD lt_Ord) |
|
493 |
thus ?case using limit |
|
494 |
by (simp add: def_transrec2 [OF Aleph_def]) |
|
495 |
qed |
|
496 |
thus ?thesis using ab . |
|
46963 | 497 |
qed |
13223 | 498 |
|
499 |
theorem Normal_Aleph: "Normal(Aleph)" |
|
76221 | 500 |
proof (rule NormalI) |
81181 | 501 |
show "i < j \<Longrightarrow> \<aleph>i < \<aleph>j" for i j |
76221 | 502 |
by (blast intro!: Aleph_increasing) |
81181 | 503 |
show "Limit(l) \<Longrightarrow> \<aleph>l = (\<Union>i<l. \<aleph>i)" for l |
76221 | 504 |
by (simp add: def_transrec2 [OF Aleph_def]) |
505 |
qed |
|
13223 | 506 |
|
507 |
end |