| author | blanchet | 
| Thu, 05 Aug 2010 14:10:18 +0200 | |
| changeset 38202 | 379fb08da97b | 
| parent 35704 | 5007843dae33 | 
| permissions | -rw-r--r-- | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1 | theory SeriesPlus | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 2 | imports Complex_Main Nat_Bijection | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 3 | begin | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 4 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 5 | text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 6 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 7 | lemma choice2: "(!!x. \<exists>y z. Q x y z) ==> \<exists>f g. \<forall>x. Q x (f x) (g x)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 8 | by metis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 9 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 10 | lemma range_subsetD: "range f \<subseteq> B \<Longrightarrow> f i \<in> B" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 11 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 12 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 13 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 14 | lemma suminf_2dimen: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 15 | fixes f:: "nat * nat \<Rightarrow> real" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 16 | assumes f_nneg: "!!m n. 0 \<le> f(m,n)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 17 | and fsums: "!!m. (\<lambda>n. f (m,n)) sums (g m)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 18 | and sumg: "summable g" | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 19 | shows "(f o prod_decode) sums suminf g" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 20 | proof (simp add: sums_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 21 |   {
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 22 | fix x | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 23 | have "0 \<le> f x" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 24 | by (cases x) (simp add: f_nneg) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 25 | } note [iff] = this | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 26 | have g_nneg: "!!m. 0 \<le> g m" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 27 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 28 | fix m | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 29 | have "0 \<le> suminf (\<lambda>n. f (m,n))" | 
| 33536 | 30 | by (rule suminf_0_le, simp add: f_nneg, metis fsums sums_iff) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 31 | thus "0 \<le> g m" using fsums [of m] | 
| 33536 | 32 | by (auto simp add: sums_iff) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 33 | qed | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 34 | show "(\<lambda>n. \<Sum>x = 0..<n. f (prod_decode x)) ----> suminf g" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 35 | proof (rule increasing_LIMSEQ, simp add: f_nneg) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 36 | fix n | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 37 |     def i \<equiv> "Max (Domain (prod_decode ` {0..<n}))"
 | 
| 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 38 |     def j \<equiv> "Max (Range (prod_decode ` {0..<n}))"
 | 
| 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 39 |     have ij: "prod_decode ` {0..<n} \<subseteq> ({0..i} \<times> {0..j})" 
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 40 | by (force simp add: i_def j_def | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 41 | intro: finite_imageI Max_ge finite_Domain finite_Range) | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 42 |     have "(\<Sum>x = 0..<n. f (prod_decode x)) = setsum f (prod_decode ` {0..<n})" 
 | 
| 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 43 |       using setsum_reindex [of prod_decode "{0..<n}" f] 
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 44 | by (simp add: o_def) | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 45 | (metis inj_prod_decode inj_prod_decode) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 46 |     also have "... \<le> setsum f ({0..i} \<times> {0..j})"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 47 | by (rule setsum_mono2) (auto simp add: ij) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 48 |     also have "... = setsum (\<lambda>m. setsum (\<lambda>n. f (m,n)) {0..j}) {0..<Suc i}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 49 | by (metis atLeast0AtMost atLeast0LessThan lessThan_Suc_atMost | 
| 33536 | 50 | setsum_cartesian_product split_eta) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 51 |     also have "... \<le> setsum g {0..<Suc i}" 
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 52 | proof (rule setsum_mono, simp add: less_Suc_eq_le) | 
| 33536 | 53 | fix m | 
| 54 | assume m: "m \<le> i" | |
| 55 | thus " (\<Sum>n = 0..j. f (m, n)) \<le> g m" using fsums [of m] | |
| 56 | by (auto simp add: sums_iff) | |
| 57 | (metis atLeastLessThanSuc_atLeastAtMost f_nneg series_pos_le f_nneg) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 58 | qed | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 59 |     finally have  "(\<Sum>x = 0..<n. f (prod_decode x)) \<le> setsum g {0..<Suc i}" .
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 60 | also have "... \<le> suminf g" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 61 | by (rule series_pos_le [OF sumg]) (simp add: g_nneg) | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 62 | finally show "(\<Sum>x = 0..<n. f (prod_decode x)) \<le> suminf g" . | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 63 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 64 | fix e :: real | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 65 | assume e: "0 < e" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 66 | with summable_sums [OF sumg] | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 67 |     obtain N where "\<bar>setsum g {0..<N} - suminf g\<bar> < e/2" and nz: "N>0"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 68 | by (simp add: sums_def LIMSEQ_iff_nz dist_real_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 69 | (metis e half_gt_zero le_refl that) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 70 |     hence gless: "suminf g < setsum g {0..<N} + e/2" using series_pos_le [OF sumg]
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 71 | by (simp add: g_nneg) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 72 |     { fix m
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 73 | assume m: "m<N" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 74 | hence enneg: "0 < e / (2 * real N)" using e | 
| 33536 | 75 | by (simp add: zero_less_divide_iff) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 76 | hence "\<exists>j. \<bar>(\<Sum>n = 0..<j. f (m, n)) - g m\<bar> < e/(2 * real N)" | 
| 33536 | 77 | using fsums [of m] m | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 78 | by (force simp add: sums_def LIMSEQ_def dist_real_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 79 |       hence "\<exists>j. g m < setsum (\<lambda>n. f (m,n)) {0..<j} + e/(2 * real N)" 
 | 
| 33536 | 80 | using fsums [of m] | 
| 81 | by (auto simp add: sums_iff) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 82 | (metis abs_diff_less_iff add_less_cancel_right eq_diff_eq') | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 83 | } | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 84 | hence "\<exists>jj. \<forall>m. m<N \<longrightarrow> g m < (\<Sum>n = 0..<jj m. f (m, n)) + e/(2 * real N)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 85 | by (force intro: choice) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 86 | then obtain jj where jj: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 87 | "!!m. m<N \<Longrightarrow> g m < (\<Sum>n = 0..<jj m. f (m, n)) + e/(2 * real N)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 88 | by auto | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 89 |     def IJ \<equiv> "SIGMA i : {0..<N}. {0..<jj i}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 90 |     have "setsum g {0..<N} <
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 91 | (\<Sum>m = 0..<N. (\<Sum>n = 0..<jj m. f (m, n)) + e/(2 * real N))" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 92 | by (auto simp add: nz jj intro: setsum_strict_mono) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 93 | also have "... = (\<Sum>m = 0..<N. \<Sum>n = 0..<jj m. f (m, n)) + e/2" using nz | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 94 | by (auto simp add: setsum_addf real_of_nat_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 95 | also have "... = setsum f IJ + e/2" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 96 | by (simp add: IJ_def setsum_Sigma) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 97 |     finally have "setsum g {0..<N} < setsum f IJ + e/2" .
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 98 | hence glessf: "suminf g < setsum f IJ + e" using gless | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 99 | by auto | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 100 | have finite_IJ: "finite IJ" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 101 | by (simp add: IJ_def) | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 102 | def NIJ \<equiv> "Max (prod_decode -` IJ)" | 
| 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 103 |     have IJsb: "IJ \<subseteq> prod_decode ` {0..NIJ}"
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 104 | proof (auto simp add: NIJ_def) | 
| 33536 | 105 | fix i j | 
| 106 | assume ij:"(i,j) \<in> IJ" | |
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 107 | hence "(i,j) = prod_decode (prod_encode (i,j))" | 
| 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 108 | by simp | 
| 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 109 |         thus "(i,j) \<in> prod_decode ` {0..Max (prod_decode -` IJ)}"
 | 
| 33536 | 110 | by (rule image_eqI) | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 111 | (simp add: ij finite_vimageI [OF finite_IJ inj_prod_decode]) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 112 | qed | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 113 |     have "setsum f IJ \<le> setsum f (prod_decode `{0..NIJ})"
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 114 | by (rule setsum_mono2) (auto simp add: IJsb) | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 115 | also have "... = (\<Sum>k = 0..NIJ. f (prod_decode k))" | 
| 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 116 | by (simp add: setsum_reindex inj_prod_decode) | 
| 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 117 | also have "... = (\<Sum>k = 0..<Suc NIJ. f (prod_decode k))" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 118 | by (metis atLeast0AtMost atLeast0LessThan lessThan_Suc_atMost) | 
| 35704 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 119 | finally have "setsum f IJ \<le> (\<Sum>k = 0..<Suc NIJ. f (prod_decode k))" . | 
| 
5007843dae33
convert HOL-Probability to use Nat_Bijection library
 huffman parents: 
33536diff
changeset | 120 | thus "\<exists>n. suminf g \<le> (\<Sum>x = 0..<n. f (prod_decode x)) + e" using glessf | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 121 | by (metis add_right_mono local.glessf not_leE order_le_less_trans less_asym) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 122 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 123 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 124 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 125 | end |