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