| author | wenzelm |
| Sun, 19 Jan 2025 21:02:03 +0100 | |
| changeset 81928 | d6366c0c9d5c |
| parent 81181 | ff2a637a449e |
| permissions | -rw-r--r-- |
| 13543 | 1 |
(* Title: ZF/Constructible/AC_in_L.thy |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
3 |
*) |
|
4 |
||
| 60770 | 5 |
section \<open>The Axiom of Choice Holds in L!\<close> |
| 13543 | 6 |
|
| 47084 | 7 |
theory AC_in_L imports Formula Separation begin |
| 13543 | 8 |
|
| 60770 | 9 |
subsection\<open>Extending a Wellordering over a List -- Lexicographic Power\<close> |
| 13543 | 10 |
|
| 60770 | 11 |
text\<open>This could be moved into a library.\<close> |
| 13543 | 12 |
|
13 |
consts |
|
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
14 |
rlist :: "[i,i]\<Rightarrow>i" |
| 13543 | 15 |
|
16 |
inductive |
|
17 |
domains "rlist(A,r)" \<subseteq> "list(A) * list(A)" |
|
18 |
intros |
|
19 |
shorterI: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
20 |
"\<lbrakk>length(l') < length(l); l' \<in> list(A); l \<in> list(A)\<rbrakk> |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
21 |
\<Longrightarrow> <l', l> \<in> rlist(A,r)" |
| 13543 | 22 |
|
23 |
sameI: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
24 |
"\<lbrakk><l',l> \<in> rlist(A,r); a \<in> A\<rbrakk> |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
25 |
\<Longrightarrow> <Cons(a,l'), Cons(a,l)> \<in> rlist(A,r)" |
| 13543 | 26 |
|
27 |
diffI: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
28 |
"\<lbrakk>length(l') = length(l); <a',a> \<in> r; |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
29 |
l' \<in> list(A); l \<in> list(A); a' \<in> A; a \<in> A\<rbrakk> |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
30 |
\<Longrightarrow> <Cons(a',l'), Cons(a,l)> \<in> rlist(A,r)" |
| 13543 | 31 |
type_intros list.intros |
32 |
||
33 |
||
| 60770 | 34 |
subsubsection\<open>Type checking\<close> |
| 13543 | 35 |
|
36 |
lemmas rlist_type = rlist.dom_subset |
|
37 |
||
38 |
lemmas field_rlist = rlist_type [THEN field_rel_subset] |
|
39 |
||
| 60770 | 40 |
subsubsection\<open>Linearity\<close> |
| 13543 | 41 |
|
42 |
lemma rlist_Nil_Cons [intro]: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
43 |
"\<lbrakk>a \<in> A; l \<in> list(A)\<rbrakk> \<Longrightarrow> <[], Cons(a,l)> \<in> rlist(A, r)" |
| 13692 | 44 |
by (simp add: shorterI) |
| 13543 | 45 |
|
46 |
lemma linear_rlist: |
|
| 47085 | 47 |
assumes r: "linear(A,r)" shows "linear(list(A),rlist(A,r))" |
48 |
proof - |
|
| 81181 | 49 |
have "xs \<in> list(A) \<Longrightarrow> ys \<in> list(A) \<Longrightarrow> \<langle>xs,ys\<rangle> \<in> rlist(A,r) \<or> xs = ys \<or> \<langle>ys,xs\<rangle> \<in> rlist(A, r)" |
50 |
for xs ys |
|
51 |
proof (induct xs arbitrary: ys rule: list.induct) |
|
52 |
case Nil |
|
53 |
thus ?case by (induct ys rule: list.induct) (auto simp add: shorterI) |
|
54 |
next |
|
55 |
case (Cons x xs) |
|
56 |
then have yConsCase: "\<langle>Cons(x,xs),Cons(y,ys)\<rangle> \<in> rlist(A,r) \<or> x=y \<and> xs = ys \<or> \<langle>Cons(y,ys), Cons(x,xs)\<rangle> \<in> rlist(A,r)" |
|
57 |
if "y \<in> A" and "ys \<in> list(A)" for y ys |
|
58 |
using that |
|
59 |
apply (rule_tac i = "length(xs)" and j = "length(ys)" in Ord_linear_lt) |
|
60 |
apply (simp_all add: shorterI) |
|
61 |
apply (rule linearE [OF r, of x y]) |
|
62 |
apply (auto simp add: diffI intro: sameI) |
|
63 |
done |
|
64 |
from \<open>ys \<in> list(A)\<close> show ?case |
|
65 |
by (cases rule: list.cases) (simp_all add: Cons rlist_Nil_Cons yConsCase) |
|
66 |
qed |
|
| 47085 | 67 |
thus ?thesis by (simp add: linear_def) |
68 |
qed |
|
| 13543 | 69 |
|
70 |
||
| 60770 | 71 |
subsubsection\<open>Well-foundedness\<close> |
| 13543 | 72 |
|
| 60770 | 73 |
text\<open>Nothing preceeds Nil in this ordering.\<close> |
| 13543 | 74 |
inductive_cases rlist_NilE: " <l,[]> \<in> rlist(A,r)" |
75 |
||
76 |
inductive_cases rlist_ConsE: " <l', Cons(x,l)> \<in> rlist(A,r)" |
|
77 |
||
78 |
lemma not_rlist_Nil [simp]: " <l,[]> \<notin> rlist(A,r)" |
|
79 |
by (blast intro: elim: rlist_NilE) |
|
80 |
||
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
81 |
lemma rlist_imp_length_le: "<l',l> \<in> rlist(A,r) \<Longrightarrow> length(l') \<le> length(l)" |
| 13543 | 82 |
apply (erule rlist.induct) |
| 13692 | 83 |
apply (simp_all add: leI) |
| 13543 | 84 |
done |
85 |
||
86 |
lemma wf_on_rlist_n: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
87 |
"\<lbrakk>n \<in> nat; wf[A](r)\<rbrakk> \<Longrightarrow> wf[{l \<in> list(A). length(l) = n}](rlist(A,r))"
|
| 13692 | 88 |
apply (induct_tac n) |
89 |
apply (rule wf_onI2, simp) |
|
90 |
apply (rule wf_onI2, clarify) |
|
91 |
apply (erule_tac a=y in list.cases, clarify) |
|
| 13543 | 92 |
apply (simp (no_asm_use)) |
| 13692 | 93 |
apply clarify |
| 13543 | 94 |
apply (simp (no_asm_use)) |
| 46823 | 95 |
apply (subgoal_tac "\<forall>l2 \<in> list(A). length(l2) = x \<longrightarrow> Cons(a,l2) \<in> B", blast) |
| 13543 | 96 |
apply (erule_tac a=a in wf_on_induct, assumption) |
97 |
apply (rule ballI) |
|
| 13692 | 98 |
apply (rule impI) |
| 13543 | 99 |
apply (erule_tac a=l2 in wf_on_induct, blast, clarify) |
| 13692 | 100 |
apply (rename_tac a' l2 l') |
101 |
apply (drule_tac x="Cons(a',l')" in bspec, typecheck) |
|
102 |
apply simp |
|
103 |
apply (erule mp, clarify) |
|
| 13543 | 104 |
apply (erule rlist_ConsE, auto) |
105 |
done |
|
106 |
||
107 |
lemma list_eq_UN_length: "list(A) = (\<Union>n\<in>nat. {l \<in> list(A). length(l) = n})"
|
|
108 |
by (blast intro: length_type) |
|
109 |
||
110 |
||
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
111 |
lemma wf_on_rlist: "wf[A](r) \<Longrightarrow> wf[list(A)](rlist(A,r))" |
| 13692 | 112 |
apply (subst list_eq_UN_length) |
113 |
apply (rule wf_on_Union) |
|
| 13543 | 114 |
apply (rule wf_imp_wf_on [OF wf_Memrel [of nat]]) |
115 |
apply (simp add: wf_on_rlist_n) |
|
| 13692 | 116 |
apply (frule rlist_type [THEN subsetD]) |
117 |
apply (simp add: length_type) |
|
| 13543 | 118 |
apply (drule rlist_imp_length_le) |
| 13692 | 119 |
apply (erule leE) |
120 |
apply (simp_all add: lt_def) |
|
| 13543 | 121 |
done |
122 |
||
123 |
||
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
124 |
lemma wf_rlist: "wf(r) \<Longrightarrow> wf(rlist(field(r),r))" |
| 13543 | 125 |
apply (simp add: wf_iff_wf_on_field) |
126 |
apply (rule wf_on_subset_A [OF _ field_rlist]) |
|
| 13692 | 127 |
apply (blast intro: wf_on_rlist) |
| 13543 | 128 |
done |
129 |
||
130 |
lemma well_ord_rlist: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
131 |
"well_ord(A,r) \<Longrightarrow> well_ord(list(A), rlist(A,r))" |
| 13543 | 132 |
apply (rule well_ordI) |
133 |
apply (simp add: well_ord_def wf_on_rlist) |
|
134 |
apply (simp add: well_ord_def tot_ord_def linear_rlist) |
|
135 |
done |
|
136 |
||
137 |
||
| 60770 | 138 |
subsection\<open>An Injection from Formulas into the Natural Numbers\<close> |
| 13543 | 139 |
|
| 69593 | 140 |
text\<open>There is a well-known bijection between \<^term>\<open>nat*nat\<close> and \<^term>\<open>nat\<close> given by the expression f(m,n) = triangle(m+n) + m, where triangle(k) |
| 13543 | 141 |
enumerates the triangular numbers and can be defined by triangle(0)=0, |
142 |
triangle(succ(k)) = succ(k + triangle(k)). Some small amount of effort is |
|
| 61798 | 143 |
needed to show that f is a bijection. We already know that such a bijection exists by the theorem \<open>well_ord_InfCard_square_eq\<close>: |
| 13692 | 144 |
@{thm[display] well_ord_InfCard_square_eq[no_vars]}
|
| 13543 | 145 |
|
| 13692 | 146 |
However, this result merely states that there is a bijection between the two |
147 |
sets. It provides no means of naming a specific bijection. Therefore, we |
|
148 |
conduct the proofs under the assumption that a bijection exists. The simplest |
|
| 60770 | 149 |
way to organize this is to use a locale.\<close> |
| 13692 | 150 |
|
| 69593 | 151 |
text\<open>Locale for any arbitrary injection between \<^term>\<open>nat*nat\<close> |
152 |
and \<^term>\<open>nat\<close>\<close> |
|
| 13543 | 153 |
locale Nat_Times_Nat = |
154 |
fixes fn |
|
155 |
assumes fn_inj: "fn \<in> inj(nat*nat, nat)" |
|
156 |
||
157 |
||
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
158 |
consts enum :: "[i,i]\<Rightarrow>i" |
| 13543 | 159 |
primrec |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
160 |
"enum(f, Member(x,y)) = f ` <0, f ` \<langle>x,y\<rangle>>" |
|
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
161 |
"enum(f, Equal(x,y)) = f ` <1, f ` \<langle>x,y\<rangle>>" |
| 13543 | 162 |
"enum(f, Nand(p,q)) = f ` <2, f ` <enum(f,p), enum(f,q)>>" |
163 |
"enum(f, Forall(p)) = f ` <succ(2), enum(f,p)>" |
|
164 |
||
165 |
lemma (in Nat_Times_Nat) fn_type [TC,simp]: |
|
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
166 |
"\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> fn`\<langle>x,y\<rangle> \<in> nat" |
| 13692 | 167 |
by (blast intro: inj_is_fun [OF fn_inj] apply_funtype) |
| 13543 | 168 |
|
169 |
lemma (in Nat_Times_Nat) fn_iff: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
170 |
"\<lbrakk>x \<in> nat; y \<in> nat; u \<in> nat; v \<in> nat\<rbrakk> |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
171 |
\<Longrightarrow> (fn`\<langle>x,y\<rangle> = fn`\<langle>u,v\<rangle>) \<longleftrightarrow> (x=u \<and> y=v)" |
| 13692 | 172 |
by (blast dest: inj_apply_equality [OF fn_inj]) |
| 13543 | 173 |
|
174 |
lemma (in Nat_Times_Nat) enum_type [TC,simp]: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
175 |
"p \<in> formula \<Longrightarrow> enum(fn,p) \<in> nat" |
| 13692 | 176 |
by (induct_tac p, simp_all) |
| 13543 | 177 |
|
178 |
lemma (in Nat_Times_Nat) enum_inject [rule_format]: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
179 |
"p \<in> formula \<Longrightarrow> \<forall>q\<in>formula. enum(fn,p) = enum(fn,q) \<longrightarrow> p=q" |
| 13692 | 180 |
apply (induct_tac p, simp_all) |
181 |
apply (rule ballI) |
|
182 |
apply (erule formula.cases) |
|
183 |
apply (simp_all add: fn_iff) |
|
184 |
apply (rule ballI) |
|
185 |
apply (erule formula.cases) |
|
186 |
apply (simp_all add: fn_iff) |
|
187 |
apply (rule ballI) |
|
188 |
apply (erule_tac a=qa in formula.cases) |
|
189 |
apply (simp_all add: fn_iff) |
|
190 |
apply blast |
|
191 |
apply (rule ballI) |
|
192 |
apply (erule_tac a=q in formula.cases) |
|
193 |
apply (simp_all add: fn_iff, blast) |
|
| 13543 | 194 |
done |
195 |
||
196 |
lemma (in Nat_Times_Nat) inj_formula_nat: |
|
197 |
"(\<lambda>p \<in> formula. enum(fn,p)) \<in> inj(formula, nat)" |
|
| 13692 | 198 |
apply (simp add: inj_def lam_type) |
199 |
apply (blast intro: enum_inject) |
|
| 13543 | 200 |
done |
201 |
||
202 |
lemma (in Nat_Times_Nat) well_ord_formula: |
|
203 |
"well_ord(formula, measure(formula, enum(fn)))" |
|
204 |
apply (rule well_ord_measure, simp) |
|
| 13692 | 205 |
apply (blast intro: enum_inject) |
| 13543 | 206 |
done |
207 |
||
208 |
lemmas nat_times_nat_lepoll_nat = |
|
209 |
InfCard_nat [THEN InfCard_square_eqpoll, THEN eqpoll_imp_lepoll] |
|
210 |
||
211 |
||
| 60770 | 212 |
text\<open>Not needed--but interesting?\<close> |
| 13543 | 213 |
theorem formula_lepoll_nat: "formula \<lesssim> nat" |
214 |
apply (insert nat_times_nat_lepoll_nat) |
|
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
215 |
unfolding lepoll_def |
| 13692 | 216 |
apply (blast intro: Nat_Times_Nat.inj_formula_nat Nat_Times_Nat.intro) |
217 |
done |
|
218 |
||
219 |
||
| 69593 | 220 |
subsection\<open>Defining the Wellordering on \<^term>\<open>DPow(A)\<close>\<close> |
| 13692 | 221 |
|
| 69593 | 222 |
text\<open>The objective is to build a wellordering on \<^term>\<open>DPow(A)\<close> from a |
223 |
given one on \<^term>\<open>A\<close>. We first introduce wellorderings for environments, |
|
224 |
which are lists built over \<^term>\<open>A\<close>. We combine it with the enumeration of |
|
| 13692 | 225 |
formulas. The order type of the resulting wellordering gives us a map from |
| 69593 | 226 |
(environment, formula) pairs into the ordinals. For each member of \<^term>\<open>DPow(A)\<close>, we take the minimum such ordinal.\<close> |
| 13692 | 227 |
|
| 21233 | 228 |
definition |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
229 |
env_form_r :: "[i,i,i]\<Rightarrow>i" where |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
230 |
\<comment> \<open>wellordering on (environment, formula) pairs\<close> |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
231 |
"env_form_r(f,r,A) \<equiv> |
| 13692 | 232 |
rmult(list(A), rlist(A, r), |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
233 |
formula, measure(formula, enum(f)))" |
| 13692 | 234 |
|
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
235 |
definition |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
236 |
env_form_map :: "[i,i,i,i]\<Rightarrow>i" where |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
237 |
\<comment> \<open>map from (environment, formula) pairs to ordinals\<close> |
| 13692 | 238 |
"env_form_map(f,r,A,z) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
239 |
\<equiv> ordermap(list(A) * formula, env_form_r(f,r,A)) ` z" |
| 13692 | 240 |
|
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
241 |
definition |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
242 |
DPow_ord :: "[i,i,i,i,i]\<Rightarrow>o" where |
| 69593 | 243 |
\<comment> \<open>predicate that holds if \<^term>\<open>k\<close> is a valid index for \<^term>\<open>X\<close>\<close> |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
244 |
"DPow_ord(f,r,A,X,k) \<equiv> |
| 13692 | 245 |
\<exists>env \<in> list(A). \<exists>p \<in> formula. |
| 76214 | 246 |
arity(p) \<le> succ(length(env)) \<and> |
247 |
X = {x\<in>A. sats(A, p, Cons(x,env))} \<and>
|
|
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
248 |
env_form_map(f,r,A,\<langle>env,p\<rangle>) = k" |
| 13692 | 249 |
|
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
250 |
definition |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
251 |
DPow_least :: "[i,i,i,i]\<Rightarrow>i" where |
| 69593 | 252 |
\<comment> \<open>function yielding the smallest index for \<^term>\<open>X\<close>\<close> |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
253 |
"DPow_least(f,r,A,X) \<equiv> \<mu> k. DPow_ord(f,r,A,X,k)" |
| 13692 | 254 |
|
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
255 |
definition |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
256 |
DPow_r :: "[i,i,i]\<Rightarrow>i" where |
| 69593 | 257 |
\<comment> \<open>a wellordering on \<^term>\<open>DPow(A)\<close>\<close> |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
258 |
"DPow_r(f,r,A) \<equiv> measure(DPow(A), DPow_least(f,r,A))" |
| 13692 | 259 |
|
260 |
||
261 |
lemma (in Nat_Times_Nat) well_ord_env_form_r: |
|
262 |
"well_ord(A,r) |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
263 |
\<Longrightarrow> well_ord(list(A) * formula, env_form_r(fn,r,A))" |
| 13692 | 264 |
by (simp add: env_form_r_def well_ord_rmult well_ord_rlist well_ord_formula) |
265 |
||
266 |
lemma (in Nat_Times_Nat) Ord_env_form_map: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
267 |
"\<lbrakk>well_ord(A,r); z \<in> list(A) * formula\<rbrakk> |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
268 |
\<Longrightarrow> Ord(env_form_map(fn,r,A,z))" |
| 13692 | 269 |
by (simp add: env_form_map_def Ord_ordermap well_ord_env_form_r) |
270 |
||
| 13702 | 271 |
lemma DPow_imp_ex_DPow_ord: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
272 |
"X \<in> DPow(A) \<Longrightarrow> \<exists>k. DPow_ord(fn,r,A,X,k)" |
| 13702 | 273 |
apply (simp add: DPow_ord_def) |
| 13692 | 274 |
apply (blast dest!: DPowD) |
275 |
done |
|
276 |
||
| 13702 | 277 |
lemma (in Nat_Times_Nat) DPow_ord_imp_Ord: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
278 |
"\<lbrakk>DPow_ord(fn,r,A,X,k); well_ord(A,r)\<rbrakk> \<Longrightarrow> Ord(k)" |
| 13702 | 279 |
apply (simp add: DPow_ord_def, clarify) |
| 13692 | 280 |
apply (simp add: Ord_env_form_map) |
| 13543 | 281 |
done |
282 |
||
| 13702 | 283 |
lemma (in Nat_Times_Nat) DPow_imp_DPow_least: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
284 |
"\<lbrakk>X \<in> DPow(A); well_ord(A,r)\<rbrakk> |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
285 |
\<Longrightarrow> DPow_ord(fn, r, A, X, DPow_least(fn,r,A,X))" |
| 13702 | 286 |
apply (simp add: DPow_least_def) |
287 |
apply (blast dest: DPow_imp_ex_DPow_ord intro: DPow_ord_imp_Ord LeastI) |
|
| 13692 | 288 |
done |
289 |
||
290 |
lemma (in Nat_Times_Nat) env_form_map_inject: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
291 |
"\<lbrakk>env_form_map(fn,r,A,u) = env_form_map(fn,r,A,v); well_ord(A,r); |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
292 |
u \<in> list(A) * formula; v \<in> list(A) * formula\<rbrakk> |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
293 |
\<Longrightarrow> u=v" |
| 13692 | 294 |
apply (simp add: env_form_map_def) |
295 |
apply (rule inj_apply_equality [OF bij_is_inj, OF ordermap_bij, |
|
296 |
OF well_ord_env_form_r], assumption+) |
|
297 |
done |
|
298 |
||
| 13702 | 299 |
lemma (in Nat_Times_Nat) DPow_ord_unique: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
300 |
"\<lbrakk>DPow_ord(fn,r,A,X,k); DPow_ord(fn,r,A,Y,k); well_ord(A,r)\<rbrakk> |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
301 |
\<Longrightarrow> X=Y" |
| 13702 | 302 |
apply (simp add: DPow_ord_def, clarify) |
| 13692 | 303 |
apply (drule env_form_map_inject, auto) |
304 |
done |
|
305 |
||
| 13702 | 306 |
lemma (in Nat_Times_Nat) well_ord_DPow_r: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
307 |
"well_ord(A,r) \<Longrightarrow> well_ord(DPow(A), DPow_r(fn,r,A))" |
| 13702 | 308 |
apply (simp add: DPow_r_def) |
| 13692 | 309 |
apply (rule well_ord_measure) |
|
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
|
310 |
apply (simp add: DPow_least_def) |
| 13702 | 311 |
apply (drule DPow_imp_DPow_least, assumption)+ |
| 13692 | 312 |
apply simp |
| 13702 | 313 |
apply (blast intro: DPow_ord_unique) |
| 13692 | 314 |
done |
315 |
||
316 |
lemma (in Nat_Times_Nat) DPow_r_type: |
|
| 13702 | 317 |
"DPow_r(fn,r,A) \<subseteq> DPow(A) * DPow(A)" |
318 |
by (simp add: DPow_r_def measure_def, blast) |
|
| 13692 | 319 |
|
| 13543 | 320 |
|
| 60770 | 321 |
subsection\<open>Limit Construction for Well-Orderings\<close> |
| 13543 | 322 |
|
| 60770 | 323 |
text\<open>Now we work towards the transfinite definition of wellorderings for |
| 69593 | 324 |
\<^term>\<open>Lset(i)\<close>. We assume as an inductive hypothesis that there is a family |
| 60770 | 325 |
of wellorderings for smaller ordinals.\<close> |
| 13543 | 326 |
|
| 21233 | 327 |
definition |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
328 |
rlimit :: "[i,i\<Rightarrow>i]\<Rightarrow>i" where |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
329 |
\<comment> \<open>Expresses the wellordering at limit ordinals. The conditional |
| 69593 | 330 |
lets us remove the premise \<^term>\<open>Limit(i)\<close> from some theorems.\<close> |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
331 |
"rlimit(i,r) \<equiv> |
| 13702 | 332 |
if Limit(i) then |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
333 |
{z: Lset(i) * Lset(i).
|
| 76214 | 334 |
\<exists>x' x. z = <x',x> \<and> |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
335 |
(lrank(x') < lrank(x) | |
| 76214 | 336 |
(lrank(x') = lrank(x) \<and> <x',x> \<in> r(succ(lrank(x)))))} |
| 13702 | 337 |
else 0" |
| 13692 | 338 |
|
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
339 |
definition |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
340 |
Lset_new :: "i\<Rightarrow>i" where |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
341 |
\<comment> \<open>This constant denotes the set of elements introduced at level |
| 69593 | 342 |
\<^term>\<open>succ(i)\<close>\<close> |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
343 |
"Lset_new(i) \<equiv> {x \<in> Lset(succ(i)). lrank(x) = i}"
|
| 13543 | 344 |
|
345 |
lemma Limit_Lset_eq2: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
346 |
"Limit(i) \<Longrightarrow> Lset(i) = (\<Union>j\<in>i. Lset_new(j))" |
| 13692 | 347 |
apply (simp add: Limit_Lset_eq) |
| 13543 | 348 |
apply (rule equalityI) |
349 |
apply safe |
|
350 |
apply (subgoal_tac "Ord(y)") |
|
351 |
prefer 2 apply (blast intro: Ord_in_Ord Limit_is_Ord) |
|
| 13692 | 352 |
apply (simp_all add: Limit_is_Ord Lset_iff_lrank_lt Lset_new_def |
353 |
Ord_mem_iff_lt) |
|
354 |
apply (blast intro: lt_trans) |
|
| 13543 | 355 |
apply (rule_tac x = "succ(lrank(x))" in bexI) |
|
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
|
356 |
apply (simp) |
| 13692 | 357 |
apply (blast intro: Limit_has_succ ltD) |
| 13543 | 358 |
done |
359 |
||
360 |
lemma wf_on_Lset: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
361 |
"wf[Lset(succ(j))](r(succ(j))) \<Longrightarrow> wf[Lset_new(j)](rlimit(i,r))" |
| 13692 | 362 |
apply (simp add: wf_on_def Lset_new_def) |
363 |
apply (erule wf_subset) |
|
| 13702 | 364 |
apply (simp add: rlimit_def, force) |
| 13543 | 365 |
done |
366 |
||
367 |
lemma wf_on_rlimit: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
368 |
"(\<forall>j<i. wf[Lset(j)](r(j))) \<Longrightarrow> wf[Lset(i)](rlimit(i,r))" |
| 13702 | 369 |
apply (case_tac "Limit(i)") |
370 |
prefer 2 |
|
371 |
apply (simp add: rlimit_def wf_on_any_0) |
|
| 13543 | 372 |
apply (simp add: Limit_Lset_eq2) |
373 |
apply (rule wf_on_Union) |
|
| 13692 | 374 |
apply (rule wf_imp_wf_on [OF wf_Memrel [of i]]) |
375 |
apply (blast intro: wf_on_Lset Limit_has_succ Limit_is_Ord ltI) |
|
| 13543 | 376 |
apply (force simp add: rlimit_def Limit_is_Ord Lset_iff_lrank_lt Lset_new_def |
377 |
Ord_mem_iff_lt) |
|
378 |
done |
|
379 |
||
380 |
lemma linear_rlimit: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
381 |
"\<lbrakk>Limit(i); \<forall>j<i. linear(Lset(j), r(j))\<rbrakk> |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
382 |
\<Longrightarrow> linear(Lset(i), rlimit(i,r))" |
| 13692 | 383 |
apply (frule Limit_is_Ord) |
384 |
apply (simp add: Limit_Lset_eq2 Lset_new_def) |
|
385 |
apply (simp add: linear_def rlimit_def Ball_def lt_Ord Lset_iff_lrank_lt) |
|
386 |
apply (simp add: ltI, clarify) |
|
387 |
apply (rename_tac u v) |
|
388 |
apply (rule_tac i="lrank(u)" and j="lrank(v)" in Ord_linear_lt, simp_all) |
|
| 46823 | 389 |
apply (drule_tac x="succ(lrank(u) \<union> lrank(v))" in ospec) |
| 13692 | 390 |
apply (simp add: ltI) |
391 |
apply (drule_tac x=u in spec, simp) |
|
392 |
apply (drule_tac x=v in spec, simp) |
|
| 13543 | 393 |
done |
394 |
||
395 |
lemma well_ord_rlimit: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
396 |
"\<lbrakk>Limit(i); \<forall>j<i. well_ord(Lset(j), r(j))\<rbrakk> |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
397 |
\<Longrightarrow> well_ord(Lset(i), rlimit(i,r))" |
| 13692 | 398 |
by (blast intro: well_ordI wf_on_rlimit well_ord_is_wf |
399 |
linear_rlimit well_ord_is_linear) |
|
| 13543 | 400 |
|
| 13702 | 401 |
lemma rlimit_cong: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
402 |
"(\<And>j. j<i \<Longrightarrow> r'(j) = r(j)) \<Longrightarrow> rlimit(i,r) = rlimit(i,r')" |
| 13702 | 403 |
apply (simp add: rlimit_def, clarify) |
404 |
apply (rule refl iff_refl Collect_cong ex_cong conj_cong)+ |
|
405 |
apply (simp add: Limit_is_Ord Lset_lrank_lt) |
|
406 |
done |
|
407 |
||
| 13543 | 408 |
|
| 69593 | 409 |
subsection\<open>Transfinite Definition of the Wellordering on \<^term>\<open>L\<close>\<close> |
| 13543 | 410 |
|
| 21233 | 411 |
definition |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
412 |
L_r :: "[i, i] \<Rightarrow> i" where |
|
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
413 |
"L_r(f) \<equiv> \<lambda>i. |
| 13702 | 414 |
transrec3(i, 0, \<lambda>x r. DPow_r(f, r, Lset(x)), |
415 |
\<lambda>x r. rlimit(x, \<lambda>y. r`y))" |
|
| 13543 | 416 |
|
| 60770 | 417 |
subsubsection\<open>The Corresponding Recursion Equations\<close> |
| 13543 | 418 |
lemma [simp]: "L_r(f,0) = 0" |
| 13702 | 419 |
by (simp add: L_r_def) |
| 13543 | 420 |
|
| 13702 | 421 |
lemma [simp]: "L_r(f, succ(i)) = DPow_r(f, L_r(f,i), Lset(i))" |
422 |
by (simp add: L_r_def) |
|
| 13543 | 423 |
|
| 60770 | 424 |
text\<open>The limit case is non-trivial because of the distinction between |
425 |
object-level and meta-level abstraction.\<close> |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
426 |
lemma [simp]: "Limit(i) \<Longrightarrow> L_r(f,i) = rlimit(i, L_r(f))" |
| 13702 | 427 |
by (simp cong: rlimit_cong add: transrec3_Limit L_r_def ltD) |
| 13543 | 428 |
|
429 |
lemma (in Nat_Times_Nat) L_r_type: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
430 |
"Ord(i) \<Longrightarrow> L_r(fn,i) \<subseteq> Lset(i) * Lset(i)" |
| 46927 | 431 |
apply (induct i rule: trans_induct3) |
| 13692 | 432 |
apply (simp_all add: Lset_succ DPow_r_type well_ord_DPow_r rlimit_def |
433 |
Transset_subset_DPow [OF Transset_Lset], blast) |
|
| 13543 | 434 |
done |
435 |
||
436 |
lemma (in Nat_Times_Nat) well_ord_L_r: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
437 |
"Ord(i) \<Longrightarrow> well_ord(Lset(i), L_r(fn,i))" |
| 46927 | 438 |
apply (induct i rule: trans_induct3) |
| 13692 | 439 |
apply (simp_all add: well_ord0 Lset_succ L_r_type well_ord_DPow_r |
440 |
well_ord_rlimit ltD) |
|
| 13543 | 441 |
done |
442 |
||
443 |
lemma well_ord_L_r: |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
444 |
"Ord(i) \<Longrightarrow> \<exists>r. well_ord(Lset(i), r)" |
| 13543 | 445 |
apply (insert nat_times_nat_lepoll_nat) |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
446 |
unfolding lepoll_def |
| 13692 | 447 |
apply (blast intro: Nat_Times_Nat.well_ord_L_r Nat_Times_Nat.intro) |
| 13543 | 448 |
done |
449 |
||
450 |
||
| 60770 | 451 |
text\<open>Every constructible set is well-ordered! Therefore the Wellordering Theorem and |
|
76220
cf8f85e2a807
fixed some theory presentation issues (?)
paulson <lp15@cam.ac.uk>
parents:
76216
diff
changeset
|
452 |
the Axiom of Choice hold in \<^term>\<open>L\<close>!\<close> |
| 47072 | 453 |
theorem L_implies_AC: assumes x: "L(x)" shows "\<exists>r. well_ord(x,r)" |
454 |
using Transset_Lset x |
|
| 13543 | 455 |
apply (simp add: Transset_def L_def) |
| 13692 | 456 |
apply (blast dest!: well_ord_L_r intro: well_ord_subset) |
| 13543 | 457 |
done |
458 |
||
|
71568
1005c50b2750
prefer strict qualification (default for 'interpretation', see 461ee3e49ad3) as proposed by Pedro Sánchez Terraf;
wenzelm
parents:
71417
diff
changeset
|
459 |
interpretation L: M_basic L by (rule M_basic_L) |
| 47084 | 460 |
|
461 |
theorem "\<forall>x[L]. \<exists>r. wellordered(L,x,r)" |
|
462 |
proof |
|
463 |
fix x |
|
464 |
assume "L(x)" |
|
465 |
then obtain r where "well_ord(x,r)" |
|
466 |
by (blast dest: L_implies_AC) |
|
467 |
thus "\<exists>r. wellordered(L,x,r)" |
|
|
71568
1005c50b2750
prefer strict qualification (default for 'interpretation', see 461ee3e49ad3) as proposed by Pedro Sánchez Terraf;
wenzelm
parents:
71417
diff
changeset
|
468 |
by (blast intro: L.well_ord_imp_relativized) |
| 47084 | 469 |
qed |
470 |
||
| 69593 | 471 |
text\<open>In order to prove \<^term>\<open> \<exists>r[L]. wellordered(L,x,r)\<close>, it's necessary to know |
472 |
that \<^term>\<open>r\<close> is actually constructible. It follows from the assumption ``\<^term>\<open>V\<close> equals \<^term>\<open>L''\<close>, |
|
| 60770 | 473 |
but this reasoning doesn't appear to work in Isabelle.\<close> |
| 47072 | 474 |
|
| 13543 | 475 |
end |