# HG changeset patch # User wenzelm # Date 1468343537 -7200 # Node ID 9d4dbb7a548a9939ed903a32f43625545fb09cae # Parent b6a1047bc164328b9eeff09bdcbbdeeaa05aadc1 more standard name; diff -r b6a1047bc164 -r 9d4dbb7a548a src/HOL/Hahn_Banach/Bounds.thy --- a/src/HOL/Hahn_Banach/Bounds.thy Tue Jul 12 16:04:19 2016 +0200 +++ b/src/HOL/Hahn_Banach/Bounds.thy Tue Jul 12 19:12:17 2016 +0200 @@ -5,7 +5,7 @@ section \Bounds\ theory Bounds -imports Main "~~/src/HOL/Library/ContNotDenum" +imports Main "~~/src/HOL/Library/Continuum_Not_Denumerable" begin locale lub = diff -r b6a1047bc164 -r 9d4dbb7a548a src/HOL/Library/ContNotDenum.thy --- a/src/HOL/Library/ContNotDenum.thy Tue Jul 12 16:04:19 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,180 +0,0 @@ -(* Title: HOL/Library/ContNotDenum.thy - Author: Benjamin Porter, Monash University, NICTA, 2005 - Author: Johannes Hölzl, TU München -*) - -section \Non-denumerability of the Continuum.\ - -theory ContNotDenum -imports Complex_Main Countable_Set -begin - -subsection \Abstract\ - -text \The following document presents a proof that the Continuum is -uncountable. It is formalised in the Isabelle/Isar theorem proving -system. - -{\em Theorem:} The Continuum \\\ is not denumerable. In other -words, there does not exist a function \f: \ \ \\ such that f is -surjective. - -{\em Outline:} An elegant informal proof of this result uses Cantor's -Diagonalisation argument. The proof presented here is not this -one. First we formalise some properties of closed intervals, then we -prove the Nested Interval Property. This property relies on the -completeness of the Real numbers and is the foundation for our -argument. Informally it states that an intersection of countable -closed intervals (where each successive interval is a subset of the -last) is non-empty. We then assume a surjective function \f: \ \ \\ exists and find a real x such that x is not in the range of f -by generating a sequence of closed intervals then using the NIP.\ - -theorem real_non_denum: "\ (\f :: nat \ real. surj f)" -proof - assume "\f::nat \ real. surj f" - then obtain f :: "nat \ real" where "surj f" .. - - txt \First we construct a sequence of nested intervals, ignoring @{term "range f"}.\ - - have "\a b c::real. a < b \ (\ka kb. ka < kb \ {ka..kb} \ {a..b} \ c \ {ka..kb})" - by (auto simp add: not_le cong: conj_cong) - (metis dense le_less_linear less_linear less_trans order_refl) - then obtain i j where ij: - "a < b \ i a b c < j a b c" - "a < b \ {i a b c .. j a b c} \ {a .. b}" - "a < b \ c \ {i a b c .. j a b c}" - for a b c :: real - by metis - - define ivl where "ivl = - rec_nat (f 0 + 1, f 0 + 2) (\n x. (i (fst x) (snd x) (f n), j (fst x) (snd x) (f n)))" - define I where "I n = {fst (ivl n) .. snd (ivl n)}" for n - - have ivl[simp]: - "ivl 0 = (f 0 + 1, f 0 + 2)" - "\n. ivl (Suc n) = (i (fst (ivl n)) (snd (ivl n)) (f n), j (fst (ivl n)) (snd (ivl n)) (f n))" - unfolding ivl_def by simp_all - - txt \This is a decreasing sequence of non-empty intervals.\ - - { fix n have "fst (ivl n) < snd (ivl n)" - by (induct n) (auto intro!: ij) } - note less = this - - have "decseq I" - unfolding I_def decseq_Suc_iff ivl fst_conv snd_conv by (intro ij allI less) - - txt \Now we apply the finite intersection property of compact sets.\ - - have "I 0 \ (\i. I i) \ {}" - proof (rule compact_imp_fip_image) - fix S :: "nat set" assume fin: "finite S" - have "{} \ I (Max (insert 0 S))" - unfolding I_def using less[of "Max (insert 0 S)"] by auto - also have "I (Max (insert 0 S)) \ (\i\insert 0 S. I i)" - using fin decseqD[OF \decseq I\, of _ "Max (insert 0 S)"] by (auto simp: Max_ge_iff) - also have "(\i\insert 0 S. I i) = I 0 \ (\i\S. I i)" - by auto - finally show "I 0 \ (\i\S. I i) \ {}" - by auto - qed (auto simp: I_def) - then obtain x where "x \ I n" for n - by blast - moreover from \surj f\ obtain j where "x = f j" - by blast - ultimately have "f j \ I (Suc j)" - by blast - with ij(3)[OF less] show False - unfolding I_def ivl fst_conv snd_conv by auto -qed - -lemma uncountable_UNIV_real: "uncountable (UNIV::real set)" - using real_non_denum unfolding uncountable_def by auto - -lemma bij_betw_open_intervals: - fixes a b c d :: real - assumes "a < b" "c < d" - shows "\f. bij_betw f {a<.. a < b" -proof - assume "uncountable {a<..a < b\, of "-pi/2" "pi/2"] by auto - then show ?thesis - by (metis bij_betw_tan uncountable_bij_betw uncountable_UNIV_real) - qed -qed - -lemma uncountable_half_open_interval_1: - fixes a :: real shows "uncountable {a.. a ax\{a<.. A" -proof - - from \countable A\ have "countable (A \ {a<..a < b\ have "\ countable {a<.. {a<.. {a<.. {a<.. {a<..x. x \ {a<.. {a<.. {}" "open S" - shows "\x\S. x \ A" -proof - - obtain x where "x \ S" - using \S \ {}\ by auto - then obtain e where "0 < e" "{y. dist y x < e} \ S" - using \open S\ by (auto simp: open_dist subset_eq) - moreover have "{y. dist y x < e} = {x - e <..< x + e}" - by (auto simp: dist_real_def) - ultimately have "uncountable (S - A)" - using uncountable_open_interval[of "x - e" "x + e"] \countable A\ - by (intro uncountable_minus_countable) (auto dest: countable_subset) - then show ?thesis - unfolding uncountable_def by auto -qed - -end diff -r b6a1047bc164 -r 9d4dbb7a548a src/HOL/Library/Continuum_Not_Denumerable.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Continuum_Not_Denumerable.thy Tue Jul 12 19:12:17 2016 +0200 @@ -0,0 +1,180 @@ +(* Title: HOL/Library/Continuum_Not_Denumerable.thy + Author: Benjamin Porter, Monash University, NICTA, 2005 + Author: Johannes Hölzl, TU München +*) + +section \Non-denumerability of the Continuum.\ + +theory Continuum_Not_Denumerable +imports Complex_Main Countable_Set +begin + +subsection \Abstract\ + +text \The following document presents a proof that the Continuum is +uncountable. It is formalised in the Isabelle/Isar theorem proving +system. + +{\em Theorem:} The Continuum \\\ is not denumerable. In other +words, there does not exist a function \f: \ \ \\ such that f is +surjective. + +{\em Outline:} An elegant informal proof of this result uses Cantor's +Diagonalisation argument. The proof presented here is not this +one. First we formalise some properties of closed intervals, then we +prove the Nested Interval Property. This property relies on the +completeness of the Real numbers and is the foundation for our +argument. Informally it states that an intersection of countable +closed intervals (where each successive interval is a subset of the +last) is non-empty. We then assume a surjective function \f: \ \ \\ exists and find a real x such that x is not in the range of f +by generating a sequence of closed intervals then using the NIP.\ + +theorem real_non_denum: "\ (\f :: nat \ real. surj f)" +proof + assume "\f::nat \ real. surj f" + then obtain f :: "nat \ real" where "surj f" .. + + txt \First we construct a sequence of nested intervals, ignoring @{term "range f"}.\ + + have "\a b c::real. a < b \ (\ka kb. ka < kb \ {ka..kb} \ {a..b} \ c \ {ka..kb})" + by (auto simp add: not_le cong: conj_cong) + (metis dense le_less_linear less_linear less_trans order_refl) + then obtain i j where ij: + "a < b \ i a b c < j a b c" + "a < b \ {i a b c .. j a b c} \ {a .. b}" + "a < b \ c \ {i a b c .. j a b c}" + for a b c :: real + by metis + + define ivl where "ivl = + rec_nat (f 0 + 1, f 0 + 2) (\n x. (i (fst x) (snd x) (f n), j (fst x) (snd x) (f n)))" + define I where "I n = {fst (ivl n) .. snd (ivl n)}" for n + + have ivl[simp]: + "ivl 0 = (f 0 + 1, f 0 + 2)" + "\n. ivl (Suc n) = (i (fst (ivl n)) (snd (ivl n)) (f n), j (fst (ivl n)) (snd (ivl n)) (f n))" + unfolding ivl_def by simp_all + + txt \This is a decreasing sequence of non-empty intervals.\ + + { fix n have "fst (ivl n) < snd (ivl n)" + by (induct n) (auto intro!: ij) } + note less = this + + have "decseq I" + unfolding I_def decseq_Suc_iff ivl fst_conv snd_conv by (intro ij allI less) + + txt \Now we apply the finite intersection property of compact sets.\ + + have "I 0 \ (\i. I i) \ {}" + proof (rule compact_imp_fip_image) + fix S :: "nat set" assume fin: "finite S" + have "{} \ I (Max (insert 0 S))" + unfolding I_def using less[of "Max (insert 0 S)"] by auto + also have "I (Max (insert 0 S)) \ (\i\insert 0 S. I i)" + using fin decseqD[OF \decseq I\, of _ "Max (insert 0 S)"] by (auto simp: Max_ge_iff) + also have "(\i\insert 0 S. I i) = I 0 \ (\i\S. I i)" + by auto + finally show "I 0 \ (\i\S. I i) \ {}" + by auto + qed (auto simp: I_def) + then obtain x where "x \ I n" for n + by blast + moreover from \surj f\ obtain j where "x = f j" + by blast + ultimately have "f j \ I (Suc j)" + by blast + with ij(3)[OF less] show False + unfolding I_def ivl fst_conv snd_conv by auto +qed + +lemma uncountable_UNIV_real: "uncountable (UNIV::real set)" + using real_non_denum unfolding uncountable_def by auto + +lemma bij_betw_open_intervals: + fixes a b c d :: real + assumes "a < b" "c < d" + shows "\f. bij_betw f {a<.. a < b" +proof + assume "uncountable {a<..a < b\, of "-pi/2" "pi/2"] by auto + then show ?thesis + by (metis bij_betw_tan uncountable_bij_betw uncountable_UNIV_real) + qed +qed + +lemma uncountable_half_open_interval_1: + fixes a :: real shows "uncountable {a.. a ax\{a<.. A" +proof - + from \countable A\ have "countable (A \ {a<..a < b\ have "\ countable {a<.. {a<.. {a<.. {a<.. {a<..x. x \ {a<.. {a<.. {}" "open S" + shows "\x\S. x \ A" +proof - + obtain x where "x \ S" + using \S \ {}\ by auto + then obtain e where "0 < e" "{y. dist y x < e} \ S" + using \open S\ by (auto simp: open_dist subset_eq) + moreover have "{y. dist y x < e} = {x - e <..< x + e}" + by (auto simp: dist_real_def) + ultimately have "uncountable (S - A)" + using uncountable_open_interval[of "x - e" "x + e"] \countable A\ + by (intro uncountable_minus_countable) (auto dest: countable_subset) + then show ?thesis + unfolding uncountable_def by auto +qed + +end diff -r b6a1047bc164 -r 9d4dbb7a548a src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Tue Jul 12 16:04:19 2016 +0200 +++ b/src/HOL/Library/Library.thy Tue Jul 12 19:12:17 2016 +0200 @@ -10,7 +10,7 @@ Bourbaki_Witt_Fixpoint Char_ord Code_Test - ContNotDenum + Continuum_Not_Denumerable Convex Combine_PER Complete_Partial_Order2 diff -r b6a1047bc164 -r 9d4dbb7a548a src/HOL/Probability/Distribution_Functions.thy --- a/src/HOL/Probability/Distribution_Functions.thy Tue Jul 12 16:04:19 2016 +0200 +++ b/src/HOL/Probability/Distribution_Functions.thy Tue Jul 12 19:12:17 2016 +0200 @@ -17,7 +17,7 @@ should be somewhere else. *) theory Distribution_Functions - imports Probability_Measure "~~/src/HOL/Library/ContNotDenum" + imports Probability_Measure "~~/src/HOL/Library/Continuum_Not_Denumerable" begin lemma UN_Ioc_eq_UNIV: "(\n. { -real n <.. real n}) = UNIV" diff -r b6a1047bc164 -r 9d4dbb7a548a src/HOL/Probability/ex/Measure_Not_CCC.thy --- a/src/HOL/Probability/ex/Measure_Not_CCC.thy Tue Jul 12 16:04:19 2016 +0200 +++ b/src/HOL/Probability/ex/Measure_Not_CCC.thy Tue Jul 12 19:12:17 2016 +0200 @@ -3,7 +3,7 @@ section \The Category of Measurable Spaces is not Cartesian Closed\ theory Measure_Not_CCC - imports "~~/src/HOL/Probability/Probability" "~~/src/HOL/Library/ContNotDenum" + imports "~~/src/HOL/Probability/Probability" "~~/src/HOL/Library/Continuum_Not_Denumerable" begin text \