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