| author | wenzelm | 
| Mon, 04 Jul 2011 20:18:19 +0200 | |
| changeset 43660 | bfc0bb115fa1 | 
| parent 40077 | c8a9eaaa2f59 | 
| child 53374 | a14d2a854c02 | 
| permissions | -rw-r--r-- | 
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
25162diff
changeset | 1 | (* Title: HOL/ex/HarmonicSeries.thy | 
| 19106 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 2 | Author: Benjamin Porter, 2006 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 3 | *) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 4 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 5 | header {* Divergence of the Harmonic Series *}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 6 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 7 | theory HarmonicSeries | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 8 | imports Complex_Main | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 9 | begin | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 10 | |
| 33509 
29e4cf2c4ea3
Turned sections into subsections (better document structure).
 webertj parents: 
28952diff
changeset | 11 | subsection {* Abstract *}
 | 
| 19106 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 12 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 13 | text {* The following document presents a proof of the Divergence of
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 14 | Harmonic Series theorem formalised in the Isabelle/Isar theorem | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 15 | proving system. | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 16 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 17 | {\em Theorem:} The series $\sum_{n=1}^{\infty} \frac{1}{n}$ does not
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 18 | converge to any number. | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 19 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 20 | {\em Informal Proof:}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 21 | The informal proof is based on the following auxillary lemmas: | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 22 |   \begin{itemize}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 23 |   \item{aux: $\sum_{n=2^m-1}^{2^m} \frac{1}{n} \geq \frac{1}{2}$}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 24 |   \item{aux2: $\sum_{n=1}^{2^M} \frac{1}{n} = 1 + \sum_{m=1}^{M} \sum_{n=2^m-1}^{2^m} \frac{1}{n}$}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 25 |   \end{itemize}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 26 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 27 |   From {\em aux} and {\em aux2} we can deduce that $\sum_{n=1}^{2^M}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 28 |   \frac{1}{n} \geq 1 + \frac{M}{2}$ for all $M$.
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 29 |   Now for contradiction, assume that $\sum_{n=1}^{\infty} \frac{1}{n}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 30 |   = s$ for some $s$. Because $\forall n. \frac{1}{n} > 0$ all the
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 31 | partial sums in the series must be less than $s$. However with our | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 32 | deduction above we can choose $N > 2*s - 2$ and thus | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 33 |   $\sum_{n=1}^{2^N} \frac{1}{n} > s$. This leads to a contradiction
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 34 |   and hence $\sum_{n=1}^{\infty} \frac{1}{n}$ is not summable.
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 35 | QED. | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 36 | *} | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 37 | |
| 33509 
29e4cf2c4ea3
Turned sections into subsections (better document structure).
 webertj parents: 
28952diff
changeset | 38 | subsection {* Formal Proof *}
 | 
| 19106 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 39 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 40 | lemma two_pow_sub: | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 41 | "0 < m \<Longrightarrow> (2::nat)^m - 2^(m - 1) = 2^(m - 1)" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 42 | by (induct m) auto | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 43 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 44 | text {* We first prove the following auxillary lemma. This lemma
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 45 | simply states that the finite sums: $\frac{1}{2}$, $\frac{1}{3} +
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 46 | \frac{1}{4}$, $\frac{1}{5} + \frac{1}{6} + \frac{1}{7} + \frac{1}{8}$
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 47 | etc. are all greater than or equal to $\frac{1}{2}$. We do this by
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 48 | observing that each term in the sum is greater than or equal to the | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 49 | last term, e.g. $\frac{1}{3} > \frac{1}{4}$ and thus $\frac{1}{3} +
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 50 | \frac{1}{4} > \frac{1}{4} + \frac{1}{4} = \frac{1}{2}$. *}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 51 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 52 | lemma harmonic_aux: | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 53 |   "\<forall>m>0. (\<Sum>n\<in>{(2::nat)^(m - 1)+1..2^m}. 1/real n) \<ge> 1/2"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 54 | (is "\<forall>m>0. (\<Sum>n\<in>(?S m). 1/real n) \<ge> 1/2") | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 55 | proof | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 56 | fix m::nat | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 57 | obtain tm where tmdef: "tm = (2::nat)^m" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 58 |   {
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 59 | assume mgt0: "0 < m" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 60 | have "\<And>x. x\<in>(?S m) \<Longrightarrow> 1/(real x) \<ge> 1/(real tm)" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 61 | proof - | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 62 | fix x::nat | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 63 | assume xs: "x\<in>(?S m)" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 64 | have xgt0: "x>0" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 65 | proof - | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 66 | from xs have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 67 | "x \<ge> 2^(m - 1) + 1" by auto | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 68 | moreover with mgt0 have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 69 | "2^(m - 1) + 1 \<ge> (1::nat)" by auto | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 70 | ultimately have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 71 | "x \<ge> 1" by (rule xtrans) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 72 | thus ?thesis by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 73 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 74 | moreover from xs have "x \<le> 2^m" by auto | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 75 | ultimately have | 
| 35345 | 76 | "inverse (real x) \<ge> inverse (real ((2::nat)^m))" | 
| 77 | by (simp del: real_of_nat_power) | |
| 19106 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 78 | moreover | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 79 | from xgt0 have "real x \<noteq> 0" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 80 | then have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 81 | "inverse (real x) = 1 / (real x)" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 82 | by (rule nonzero_inverse_eq_divide) | 
| 25162 | 83 | moreover from mgt0 have "real tm \<noteq> 0" by (simp add: tmdef) | 
| 19106 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 84 | then have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 85 | "inverse (real tm) = 1 / (real tm)" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 86 | by (rule nonzero_inverse_eq_divide) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 87 | ultimately show | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 88 | "1/(real x) \<ge> 1/(real tm)" by (auto simp add: tmdef) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 89 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 90 | then have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 91 | "(\<Sum>n\<in>(?S m). 1 / real n) \<ge> (\<Sum>n\<in>(?S m). 1/(real tm))" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 92 | by (rule setsum_mono) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 93 | moreover have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 94 | "(\<Sum>n\<in>(?S m). 1/(real tm)) = 1/2" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 95 | proof - | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 96 | have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 97 | "(\<Sum>n\<in>(?S m). 1/(real tm)) = | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 98 | (1/(real tm))*(\<Sum>n\<in>(?S m). 1)" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 99 | by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 100 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 101 | "\<dots> = ((1/(real tm)) * real (card (?S m)))" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 102 | by (simp add: real_of_card real_of_nat_def) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 103 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 104 | "\<dots> = ((1/(real tm)) * real (tm - (2^(m - 1))))" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 105 | by (simp add: tmdef) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 106 | also from mgt0 have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 107 | "\<dots> = ((1/(real tm)) * real ((2::nat)^(m - 1)))" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 108 | by (auto simp: tmdef dest: two_pow_sub) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 109 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 110 | "\<dots> = (real (2::nat))^(m - 1) / (real (2::nat))^m" | 
| 35345 | 111 | by (simp add: tmdef) | 
| 19106 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 112 | also from mgt0 have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 113 | "\<dots> = (real (2::nat))^(m - 1) / (real (2::nat))^((m - 1) + 1)" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 114 | by auto | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 115 | also have "\<dots> = 1/2" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 116 | finally show ?thesis . | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 117 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 118 | ultimately have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 119 | "(\<Sum>n\<in>(?S m). 1 / real n) \<ge> 1/2" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 120 | by - (erule subst) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 121 | } | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 122 | thus "0 < m \<longrightarrow> 1 / 2 \<le> (\<Sum>n\<in>(?S m). 1 / real n)" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 123 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 124 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 125 | text {* We then show that the sum of a finite number of terms from the
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 126 | harmonic series can be regrouped in increasing powers of 2. For | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 127 | example: $1 + \frac{1}{2} + \frac{1}{3} + \frac{1}{4} + \frac{1}{5} +
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 128 | \frac{1}{6} + \frac{1}{7} + \frac{1}{8} = 1 + (\frac{1}{2}) +
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 129 | (\frac{1}{3} + \frac{1}{4}) + (\frac{1}{5} + \frac{1}{6} + \frac{1}{7}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 130 | + \frac{1}{8})$. *}
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 131 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 132 | lemma harmonic_aux2 [rule_format]: | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 133 |   "0<M \<Longrightarrow> (\<Sum>n\<in>{1..(2::nat)^M}. 1/real n) =
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 134 |    (1 + (\<Sum>m\<in>{1..M}. \<Sum>n\<in>{(2::nat)^(m - 1)+1..2^m}. 1/real n))"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 135 | (is "0<M \<Longrightarrow> ?LHS M = ?RHS M") | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 136 | proof (induct M) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 137 | case 0 show ?case by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 138 | next | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 139 | case (Suc M) | 
| 23464 | 140 | have ant: "0 < Suc M" by fact | 
| 19106 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 141 |   {
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 142 | have suc: "?LHS (Suc M) = ?RHS (Suc M)" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 143 | proof cases -- "show that LHS = c and RHS = c, and thus LHS = RHS" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 144 | assume mz: "M=0" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 145 |       {
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 146 | then have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 147 | "?LHS (Suc M) = ?LHS 1" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 148 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 149 |           "\<dots> = (\<Sum>n\<in>{(1::nat)..2}. 1/real n)" by simp
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 150 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 151 |           "\<dots> = ((\<Sum>n\<in>{Suc 1..2}. 1/real n) + 1/(real (1::nat)))"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 152 | by (subst setsum_head) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 153 | (auto simp: atLeastSucAtMost_greaterThanAtMost) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 154 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 155 |           "\<dots> = ((\<Sum>n\<in>{2..2::nat}. 1/real n) + 1/(real (1::nat)))"
 | 
| 40077 | 156 | by (simp add: eval_nat_numeral) | 
| 19106 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 157 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 158 | "\<dots> = 1/(real (2::nat)) + 1/(real (1::nat))" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 159 | finally have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 160 | "?LHS (Suc M) = 1/2 + 1" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 161 | } | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 162 | moreover | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 163 |       {
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 164 | from mz have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 165 | "?RHS (Suc M) = ?RHS 1" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 166 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 167 |           "\<dots> = (\<Sum>n\<in>{((2::nat)^0)+1..2^1}. 1/real n) + 1"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 168 | by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 169 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 170 |           "\<dots> = (\<Sum>n\<in>{2::nat..2}. 1/real n) + 1"
 | 
| 36846 
0f67561ed5a6
Added atLeastAtMost_singleton_iff, atLeastAtMost_singleton'
 hoelzl parents: 
35345diff
changeset | 171 | by (auto simp: atLeastAtMost_singleton') | 
| 19106 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 172 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 173 | "\<dots> = 1/2 + 1" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 174 | by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 175 | finally have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 176 | "?RHS (Suc M) = 1/2 + 1" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 177 | } | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 178 | ultimately show "?LHS (Suc M) = ?RHS (Suc M)" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 179 | next | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 180 | assume mnz: "M\<noteq>0" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 181 | then have mgtz: "M>0" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 182 | with Suc have suc: | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 183 | "(?LHS M) = (?RHS M)" by blast | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 184 | have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 185 | "(?LHS (Suc M)) = | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 186 |          ((?LHS M) + (\<Sum>n\<in>{(2::nat)^M+1..2^(Suc M)}. 1 / real n))"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 187 | proof - | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 188 | have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 189 |           "{1..(2::nat)^(Suc M)} =
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 190 |            {1..(2::nat)^M}\<union>{(2::nat)^M+1..(2::nat)^(Suc M)}"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 191 | by auto | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 192 | moreover have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 193 |           "{1..(2::nat)^M}\<inter>{(2::nat)^M+1..(2::nat)^(Suc M)} = {}"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 194 | by auto | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 195 | moreover have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 196 |           "finite {1..(2::nat)^M}" and "finite {(2::nat)^M+1..(2::nat)^(Suc M)}"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 197 | by auto | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 198 | ultimately show ?thesis | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 199 | by (auto intro: setsum_Un_disjoint) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 200 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 201 | moreover | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 202 |       {
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 203 | have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 204 | "(?RHS (Suc M)) = | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 205 |            (1 + (\<Sum>m\<in>{1..M}.  \<Sum>n\<in>{(2::nat)^(m - 1)+1..2^m}. 1/real n) +
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 206 |            (\<Sum>n\<in>{(2::nat)^(Suc M - 1)+1..2^(Suc M)}. 1/real n))" by simp
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 207 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 208 |           "\<dots> = (?RHS M) + (\<Sum>n\<in>{(2::nat)^M+1..2^(Suc M)}. 1/real n)"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 209 | by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 210 | also from suc have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 211 |           "\<dots> = (?LHS M) +  (\<Sum>n\<in>{(2::nat)^M+1..2^(Suc M)}. 1/real n)"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 212 | by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 213 | finally have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 214 | "(?RHS (Suc M)) = \<dots>" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 215 | } | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 216 | ultimately show "?LHS (Suc M) = ?RHS (Suc M)" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 217 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 218 | } | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 219 | thus ?case by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 220 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 221 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 222 | text {* Using @{thm [source] harmonic_aux} and @{thm [source] harmonic_aux2} we now show
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 223 | that each group sum is greater than or equal to $\frac{1}{2}$ and thus
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 224 | the finite sum is bounded below by a value proportional to the number | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 225 | of elements we choose. *} | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 226 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 227 | lemma harmonic_aux3 [rule_format]: | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 228 |   shows "\<forall>(M::nat). (\<Sum>n\<in>{1..(2::nat)^M}. 1 / real n) \<ge> 1 + (real M)/2"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 229 | (is "\<forall>M. ?P M \<ge> _") | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 230 | proof (rule allI, cases) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 231 | fix M::nat | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 232 | assume "M=0" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 233 | then show "?P M \<ge> 1 + (real M)/2" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 234 | next | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 235 | fix M::nat | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 236 | assume "M\<noteq>0" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 237 | then have "M > 0" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 238 | then have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 239 | "(?P M) = | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 240 |      (1 + (\<Sum>m\<in>{1..M}. \<Sum>n\<in>{(2::nat)^(m - 1)+1..2^m}. 1/real n))"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 241 | by (rule harmonic_aux2) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 242 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 243 |     "\<dots> \<ge> (1 + (\<Sum>m\<in>{1..M}. 1/2))"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 244 | proof - | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 245 | let ?f = "(\<lambda>x. 1/2)" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 246 |     let ?g = "(\<lambda>x. (\<Sum>n\<in>{(2::nat)^(x - 1)+1..2^x}. 1/real n))"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 247 |     from harmonic_aux have "\<And>x. x\<in>{1..M} \<Longrightarrow> ?f x \<le> ?g x" by simp
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 248 |     then have "(\<Sum>m\<in>{1..M}. ?g m) \<ge> (\<Sum>m\<in>{1..M}. ?f m)" by (rule setsum_mono)
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 249 | thus ?thesis by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 250 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 251 |   finally have "(?P M) \<ge> (1 + (\<Sum>m\<in>{1..M}. 1/2))" .
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 252 | moreover | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 253 |   {
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 254 | have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 255 |       "(\<Sum>m\<in>{1..M}. (1::real)/2) = 1/2 * (\<Sum>m\<in>{1..M}. 1)"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 256 | by auto | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 257 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 258 |       "\<dots> = 1/2*(real (card {1..M}))"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 259 | by (simp only: real_of_card[symmetric]) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 260 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 261 | "\<dots> = 1/2*(real M)" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 262 | also have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 263 | "\<dots> = (real M)/2" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 264 |     finally have "(\<Sum>m\<in>{1..M}. (1::real)/2) = (real M)/2" .
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 265 | } | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 266 | ultimately show "(?P M) \<ge> (1 + (real M)/2)" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 267 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 268 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 269 | text {* The final theorem shows that as we take more and more elements
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 270 | (see @{thm [source] harmonic_aux3}) we get an ever increasing sum. By assuming
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 271 | the sum converges, the lemma @{thm [source] series_pos_less} ( @{thm
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 272 | series_pos_less} ) states that each sum is bounded above by the | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 273 | series' limit. This contradicts our first statement and thus we prove | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 274 | that the harmonic series is divergent. *} | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 275 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 276 | theorem DivergenceOfHarmonicSeries: | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 277 | shows "\<not>summable (\<lambda>n. 1/real (Suc n))" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 278 | (is "\<not>summable ?f") | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 279 | proof -- "by contradiction" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 280 | let ?s = "suminf ?f" -- "let ?s equal the sum of the harmonic series" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 281 | assume sf: "summable ?f" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 282 | then obtain n::nat where ndef: "n = nat \<lceil>2 * ?s\<rceil>" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 283 | then have ngt: "1 + real n/2 > ?s" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 284 | proof - | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 285 | have "\<forall>n. 0 \<le> ?f n" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 286 | with sf have "?s \<ge> 0" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 287 | by - (rule suminf_0_le, simp_all) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 288 | then have cgt0: "\<lceil>2*?s\<rceil> \<ge> 0" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 289 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 290 | from ndef have "n = nat \<lceil>(2*?s)\<rceil>" . | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 291 | then have "real n = real (nat \<lceil>2*?s\<rceil>)" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 292 | with cgt0 have "real n = real \<lceil>2*?s\<rceil>" | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 293 | by (auto dest: real_nat_eq_real) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 294 | then have "real n \<ge> 2*(?s)" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 295 | then have "real n/2 \<ge> (?s)" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 296 | then show "1 + real n/2 > (?s)" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 297 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 298 | |
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 299 | obtain j where jdef: "j = (2::nat)^n" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 300 | have "\<forall>m\<ge>j. 0 < ?f m" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 301 |   with sf have "(\<Sum>i\<in>{0..<j}. ?f i) < ?s" by (rule series_pos_less)
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 302 |   then have "(\<Sum>i\<in>{1..<Suc j}. 1/(real i)) < ?s"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 303 | apply - | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 304 | apply (subst(asm) setsum_shift_bounds_Suc_ivl [symmetric]) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 305 | by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 306 | with jdef have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 307 |     "(\<Sum>i\<in>{1..< Suc ((2::nat)^n)}. 1 / (real i)) < ?s" by simp
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 308 | then have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 309 |     "(\<Sum>i\<in>{1..(2::nat)^n}. 1 / (real i)) < ?s"
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 310 | by (simp only: atLeastLessThanSuc_atLeastAtMost) | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 311 | moreover from harmonic_aux3 have | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 312 |     "(\<Sum>i\<in>{1..(2::nat)^n}. 1 / (real i)) \<ge> 1 + real n/2" by simp
 | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 313 | moreover from ngt have "1 + real n/2 > ?s" by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 314 | ultimately show False by simp | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 315 | qed | 
| 
6e6b5b1fdc06
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
 kleing parents: diff
changeset | 316 | |
| 35345 | 317 | end |