| author | huffman | 
| Tue, 01 Jul 2008 03:14:00 +0200 | |
| changeset 27416 | 07e04ab0177a | 
| parent 24893 | b8ef7afe3a6b | 
| child 27678 | 85ea2be46c71 | 
| permissions | -rw-r--r-- | 
| 5464 | 1 | (* Title: ZF/AC/WO1_WO7.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, CU Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | WO7 <-> LEMMA <-> WO1 (Rubin & Rubin p. 5) | |
| 7 | LEMMA is the sentence denoted by (**) | |
| 12776 | 8 | |
| 9 | Also, WO1 <-> WO8 | |
| 5464 | 10 | *) | 
| 11 | ||
| 16417 | 12 | theory WO1_WO7 imports AC_Equiv begin | 
| 5464 | 13 | |
| 24893 | 14 | definition | 
| 5464 | 15 | "LEMMA == | 
| 12776 | 16 | \<forall>X. ~Finite(X) --> (\<exists>R. well_ord(X,R) & ~well_ord(X,converse(R)))" | 
| 17 | ||
| 18 | (* ********************************************************************** *) | |
| 19 | (* It is easy to see that WO7 is equivalent to (**) *) | |
| 20 | (* ********************************************************************** *) | |
| 21 | ||
| 22 | lemma WO7_iff_LEMMA: "WO7 <-> LEMMA" | |
| 23 | apply (unfold WO7_def LEMMA_def) | |
| 24 | apply (blast intro: Finite_well_ord_converse) | |
| 25 | done | |
| 26 | ||
| 27 | (* ********************************************************************** *) | |
| 28 | (* It is also easy to show that LEMMA implies WO1. *) | |
| 29 | (* ********************************************************************** *) | |
| 30 | ||
| 31 | lemma LEMMA_imp_WO1: "LEMMA ==> WO1" | |
| 32 | apply (unfold WO1_def LEMMA_def Finite_def eqpoll_def) | |
| 33 | apply (blast intro!: well_ord_rvimage [OF bij_is_inj nat_implies_well_ord]) | |
| 34 | done | |
| 35 | ||
| 36 | (* ********************************************************************** *) | |
| 37 | (* The Rubins' proof of the other implication is contained within the *) | |
| 38 | (* following sentence \<in> *) | |
| 39 | (* "... each infinite ordinal is well ordered by < but not by >." *) | |
| 40 | (* This statement can be proved by the following two theorems. *) | |
| 41 | (* But moreover we need to show similar property for any well ordered *) | |
| 42 | (* infinite set. It is not very difficult thanks to Isabelle order types *) | |
| 43 | (* We show that if a set is well ordered by some relation and by its *) | |
| 44 | (* converse, then apropriate order type is well ordered by the converse *) | |
| 45 | (* of it's membership relation, which in connection with the previous *) | |
| 46 | (* gives the conclusion. *) | |
| 47 | (* ********************************************************************** *) | |
| 48 | ||
| 49 | lemma converse_Memrel_not_wf_on: | |
| 50 | "[| Ord(a); ~Finite(a) |] ==> ~wf[a](converse(Memrel(a)))" | |
| 51 | apply (unfold wf_on_def wf_def) | |
| 12820 | 52 | apply (drule nat_le_infinite_Ord [THEN le_imp_subset], assumption) | 
| 12776 | 53 | apply (rule notI) | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12820diff
changeset | 54 | apply (erule_tac x = nat in allE, blast) | 
| 12776 | 55 | done | 
| 56 | ||
| 57 | lemma converse_Memrel_not_well_ord: | |
| 58 | "[| Ord(a); ~Finite(a) |] ==> ~well_ord(a,converse(Memrel(a)))" | |
| 59 | apply (unfold well_ord_def) | |
| 60 | apply (blast dest: converse_Memrel_not_wf_on) | |
| 61 | done | |
| 62 | ||
| 63 | lemma well_ord_rvimage_ordertype: | |
| 64 | "well_ord(A,r) \<Longrightarrow> | |
| 65 | rvimage (ordertype(A,r), converse(ordermap(A,r)),r) = | |
| 66 | Memrel(ordertype(A,r))" | |
| 67 | by (blast intro: ordertype_ord_iso [THEN ord_iso_sym] ord_iso_rvimage_eq | |
| 68 | Memrel_type [THEN subset_Int_iff [THEN iffD1]] trans) | |
| 69 | ||
| 70 | lemma well_ord_converse_Memrel: | |
| 71 | "[| well_ord(A,r); well_ord(A,converse(r)) |] | |
| 72 | ==> well_ord(ordertype(A,r), converse(Memrel(ordertype(A,r))))" | |
| 73 | apply (subst well_ord_rvimage_ordertype [symmetric], assumption) | |
| 74 | apply (rule rvimage_converse [THEN subst]) | |
| 75 | apply (blast intro: ordertype_ord_iso ord_iso_sym ord_iso_is_bij | |
| 76 | bij_is_inj well_ord_rvimage) | |
| 77 | done | |
| 78 | ||
| 79 | lemma WO1_imp_LEMMA: "WO1 ==> LEMMA" | |
| 80 | apply (unfold WO1_def LEMMA_def, clarify) | |
| 81 | apply (blast dest: well_ord_converse_Memrel | |
| 82 | Ord_ordertype [THEN converse_Memrel_not_well_ord] | |
| 83 | intro: ordertype_ord_iso ord_iso_is_bij bij_is_inj lepoll_Finite | |
| 84 | lepoll_def [THEN def_imp_iff, THEN iffD2] ) | |
| 85 | done | |
| 86 | ||
| 87 | lemma WO1_iff_WO7: "WO1 <-> WO7" | |
| 88 | apply (simp add: WO7_iff_LEMMA) | |
| 89 | apply (blast intro: LEMMA_imp_WO1 WO1_imp_LEMMA) | |
| 90 | done | |
| 91 | ||
| 92 | ||
| 93 | ||
| 94 | (* ********************************************************************** *) | |
| 95 | (* The proof of WO8 <-> WO1 (Rubin & Rubin p. 6) *) | |
| 96 | (* ********************************************************************** *) | |
| 97 | ||
| 98 | lemma WO1_WO8: "WO1 ==> WO8" | |
| 99 | by (unfold WO1_def WO8_def, fast) | |
| 100 | ||
| 101 | ||
| 102 | (* The implication "WO8 ==> WO1": a faithful image of Rubin & Rubin's proof*) | |
| 103 | lemma WO8_WO1: "WO8 ==> WO1" | |
| 104 | apply (unfold WO1_def WO8_def) | |
| 105 | apply (rule allI) | |
| 106 | apply (erule_tac x = "{{x}. x \<in> A}" in allE)
 | |
| 107 | apply (erule impE) | |
| 108 |  apply (rule_tac x = "\<lambda>a \<in> {{x}. x \<in> A}. THE x. a={x}" in exI)
 | |
| 109 | apply (force intro!: lam_type simp add: singleton_eq_iff the_equality) | |
| 110 | apply (blast intro: lam_sing_bij bij_is_inj well_ord_rvimage) | |
| 111 | done | |
| 5464 | 112 | |
| 113 | end |