| author | wenzelm | 
| Mon, 22 Jan 2024 14:40:30 +0100 | |
| changeset 79513 | 292605271dcf | 
| parent 75866 | 9eeed5c424f9 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 62479 | 1 | (* Title: HOL/Nonstandard_Analysis/HSeries.thy | 
| 2 | Author: Jacques D. Fleuriot | |
| 3 | Copyright: 1998 University of Cambridge | |
| 27468 | 4 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 5 | Converted to Isar and polished by lcp | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 6 | *) | 
| 27468 | 7 | |
| 64604 | 8 | section \<open>Finite Summation and Infinite Series for Hyperreals\<close> | 
| 27468 | 9 | |
| 10 | theory HSeries | |
| 64604 | 11 | imports HSEQ | 
| 27468 | 12 | begin | 
| 13 | ||
| 64604 | 14 | definition sumhr :: "hypnat \<times> hypnat \<times> (nat \<Rightarrow> real) \<Rightarrow> hypreal" | 
| 15 |   where "sumhr = (\<lambda>(M,N,f). starfun2 (\<lambda>m n. sum f {m..<n}) M N)"
 | |
| 16 | ||
| 17 | definition NSsums :: "(nat \<Rightarrow> real) \<Rightarrow> real \<Rightarrow> bool" (infixr "NSsums" 80) | |
| 18 |   where "f NSsums s = (\<lambda>n. sum f {..<n}) \<longlonglongrightarrow>\<^sub>N\<^sub>S s"
 | |
| 27468 | 19 | |
| 64604 | 20 | definition NSsummable :: "(nat \<Rightarrow> real) \<Rightarrow> bool" | 
| 21 | where "NSsummable f \<longleftrightarrow> (\<exists>s. f NSsums s)" | |
| 27468 | 22 | |
| 64604 | 23 | definition NSsuminf :: "(nat \<Rightarrow> real) \<Rightarrow> real" | 
| 24 | where "NSsuminf f = (THE s. f NSsums s)" | |
| 27468 | 25 | |
| 64604 | 26 | lemma sumhr_app: "sumhr (M, N, f) = ( *f2* (\<lambda>m n. sum f {m..<n})) M N"
 | 
| 27 | by (simp add: sumhr_def) | |
| 27468 | 28 | |
| 69597 | 29 | text \<open>Base case in definition of \<^term>\<open>sumr\<close>.\<close> | 
| 64604 | 30 | lemma sumhr_zero [simp]: "\<And>m. sumhr (m, 0, f) = 0" | 
| 31 | unfolding sumhr_app by transfer simp | |
| 27468 | 32 | |
| 69597 | 33 | text \<open>Recursive case in definition of \<^term>\<open>sumr\<close>.\<close> | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 34 | lemma sumhr_if: | 
| 64604 | 35 | "\<And>m n. sumhr (m, n + 1, f) = (if n + 1 \<le> m then 0 else sumhr (m, n, f) + ( *f* f) n)" | 
| 36 | unfolding sumhr_app by transfer simp | |
| 37 | ||
| 38 | lemma sumhr_Suc_zero [simp]: "\<And>n. sumhr (n + 1, n, f) = 0" | |
| 39 | unfolding sumhr_app by transfer simp | |
| 27468 | 40 | |
| 64604 | 41 | lemma sumhr_eq_bounds [simp]: "\<And>n. sumhr (n, n, f) = 0" | 
| 42 | unfolding sumhr_app by transfer simp | |
| 27468 | 43 | |
| 64604 | 44 | lemma sumhr_Suc [simp]: "\<And>m. sumhr (m, m + 1, f) = ( *f* f) m" | 
| 45 | unfolding sumhr_app by transfer simp | |
| 27468 | 46 | |
| 64604 | 47 | lemma sumhr_add_lbound_zero [simp]: "\<And>k m. sumhr (m + k, k, f) = 0" | 
| 48 | unfolding sumhr_app by transfer simp | |
| 27468 | 49 | |
| 64604 | 50 | lemma sumhr_add: "\<And>m n. sumhr (m, n, f) + sumhr (m, n, g) = sumhr (m, n, \<lambda>i. f i + g i)" | 
| 51 | unfolding sumhr_app by transfer (rule sum.distrib [symmetric]) | |
| 27468 | 52 | |
| 64604 | 53 | lemma sumhr_mult: "\<And>m n. hypreal_of_real r * sumhr (m, n, f) = sumhr (m, n, \<lambda>n. r * f n)" | 
| 54 | unfolding sumhr_app by transfer (rule sum_distrib_left) | |
| 27468 | 55 | |
| 64604 | 56 | lemma sumhr_split_add: "\<And>n p. n < p \<Longrightarrow> sumhr (0, n, f) + sumhr (n, p, f) = sumhr (0, p, f)" | 
| 70097 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 57 | unfolding sumhr_app by transfer (simp add: sum.atLeastLessThan_concat) | 
| 27468 | 58 | |
| 64604 | 59 | lemma sumhr_split_diff: "n < p \<Longrightarrow> sumhr (0, p, f) - sumhr (0, n, f) = sumhr (n, p, f)" | 
| 60 | by (drule sumhr_split_add [symmetric, where f = f]) simp | |
| 27468 | 61 | |
| 64604 | 62 | lemma sumhr_hrabs: "\<And>m n. \<bar>sumhr (m, n, f)\<bar> \<le> sumhr (m, n, \<lambda>i. \<bar>f i\<bar>)" | 
| 63 | unfolding sumhr_app by transfer (rule sum_abs) | |
| 27468 | 64 | |
| 64604 | 65 | text \<open>Other general version also needed.\<close> | 
| 27468 | 66 | lemma sumhr_fun_hypnat_eq: | 
| 64604 | 67 | "(\<forall>r. m \<le> r \<and> r < n \<longrightarrow> f r = g r) \<longrightarrow> | 
| 68 | sumhr (hypnat_of_nat m, hypnat_of_nat n, f) = | |
| 69 | sumhr (hypnat_of_nat m, hypnat_of_nat n, g)" | |
| 70 | unfolding sumhr_app by transfer simp | |
| 27468 | 71 | |
| 64604 | 72 | lemma sumhr_const: "\<And>n. sumhr (0, n, \<lambda>i. r) = hypreal_of_hypnat n * hypreal_of_real r" | 
| 73 | unfolding sumhr_app by transfer simp | |
| 27468 | 74 | |
| 64604 | 75 | lemma sumhr_less_bounds_zero [simp]: "\<And>m n. n < m \<Longrightarrow> sumhr (m, n, f) = 0" | 
| 76 | unfolding sumhr_app by transfer simp | |
| 27468 | 77 | |
| 64604 | 78 | lemma sumhr_minus: "\<And>m n. sumhr (m, n, \<lambda>i. - f i) = - sumhr (m, n, f)" | 
| 79 | unfolding sumhr_app by transfer (rule sum_negf) | |
| 27468 | 80 | |
| 81 | lemma sumhr_shift_bounds: | |
| 64604 | 82 | "\<And>m n. sumhr (m + hypnat_of_nat k, n + hypnat_of_nat k, f) = | 
| 83 | sumhr (m, n, \<lambda>i. f (i + k))" | |
| 70113 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
70097diff
changeset | 84 | unfolding sumhr_app by transfer (rule sum.shift_bounds_nat_ivl) | 
| 27468 | 85 | |
| 86 | ||
| 64604 | 87 | subsection \<open>Nonstandard Sums\<close> | 
| 27468 | 88 | |
| 64604 | 89 | text \<open>Infinite sums are obtained by summing to some infinite hypernatural | 
| 69597 | 90 | (such as \<^term>\<open>whn\<close>).\<close> | 
| 64604 | 91 | lemma sumhr_hypreal_of_hypnat_omega: "sumhr (0, whn, \<lambda>i. 1) = hypreal_of_hypnat whn" | 
| 92 | by (simp add: sumhr_const) | |
| 27468 | 93 | |
| 75866 | 94 | lemma whn_eq_\<omega>m1: "hypreal_of_hypnat whn = \<omega> - 1" | 
| 95 | unfolding star_class_defs omega_def hypnat_omega_def of_hypnat_def star_of_def | |
| 96 | by (simp add: starfun_star_n starfun2_star_n) | |
| 97 | ||
| 64604 | 98 | lemma sumhr_hypreal_omega_minus_one: "sumhr(0, whn, \<lambda>i. 1) = \<omega> - 1" | 
| 75866 | 99 | by (simp add: sumhr_const whn_eq_\<omega>m1) | 
| 27468 | 100 | |
| 64604 | 101 | lemma sumhr_minus_one_realpow_zero [simp]: "\<And>N. sumhr (0, N + N, \<lambda>i. (-1) ^ (i + 1)) = 0" | 
| 102 | unfolding sumhr_app | |
| 75866 | 103 | by transfer (induct_tac N, auto) | 
| 27468 | 104 | |
| 105 | lemma sumhr_interval_const: | |
| 64604 | 106 | "(\<forall>n. m \<le> Suc n \<longrightarrow> f n = r) \<and> m \<le> na \<Longrightarrow> | 
| 107 | sumhr (hypnat_of_nat m, hypnat_of_nat na, f) = hypreal_of_nat (na - m) * hypreal_of_real r" | |
| 108 | unfolding sumhr_app by transfer simp | |
| 27468 | 109 | |
| 64604 | 110 | lemma starfunNat_sumr: "\<And>N. ( *f* (\<lambda>n. sum f {0..<n})) N = sumhr (0, N, f)"
 | 
| 111 | unfolding sumhr_app by transfer (rule refl) | |
| 27468 | 112 | |
| 64604 | 113 | lemma sumhr_hrabs_approx [simp]: "sumhr (0, M, f) \<approx> sumhr (0, N, f) \<Longrightarrow> \<bar>sumhr (M, N, f)\<bar> \<approx> 0" | 
| 114 | using linorder_less_linear [where x = M and y = N] | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
64604diff
changeset | 115 | by (metis (no_types, lifting) abs_zero approx_hrabs approx_minus_iff approx_refl approx_sym sumhr_eq_bounds sumhr_less_bounds_zero sumhr_split_diff) | 
| 64604 | 116 | |
| 117 | ||
| 118 | subsection \<open>Infinite sums: Standard and NS theorems\<close> | |
| 27468 | 119 | |
| 64604 | 120 | lemma sums_NSsums_iff: "f sums l \<longleftrightarrow> f NSsums l" | 
| 121 | by (simp add: sums_def NSsums_def LIMSEQ_NSLIMSEQ_iff) | |
| 27468 | 122 | |
| 64604 | 123 | lemma summable_NSsummable_iff: "summable f \<longleftrightarrow> NSsummable f" | 
| 124 | by (simp add: summable_def NSsummable_def sums_NSsums_iff) | |
| 27468 | 125 | |
| 64604 | 126 | lemma suminf_NSsuminf_iff: "suminf f = NSsuminf f" | 
| 127 | by (simp add: suminf_def NSsuminf_def sums_NSsums_iff) | |
| 27468 | 128 | |
| 64604 | 129 | lemma NSsums_NSsummable: "f NSsums l \<Longrightarrow> NSsummable f" | 
| 130 | unfolding NSsums_def NSsummable_def by blast | |
| 27468 | 131 | |
| 64604 | 132 | lemma NSsummable_NSsums: "NSsummable f \<Longrightarrow> f NSsums (NSsuminf f)" | 
| 133 | unfolding NSsummable_def NSsuminf_def NSsums_def | |
| 134 | by (blast intro: theI NSLIMSEQ_unique) | |
| 27468 | 135 | |
| 64604 | 136 | lemma NSsums_unique: "f NSsums s \<Longrightarrow> s = NSsuminf f" | 
| 137 | by (simp add: suminf_NSsuminf_iff [symmetric] sums_NSsums_iff sums_unique) | |
| 27468 | 138 | |
| 64604 | 139 | lemma NSseries_zero: "\<forall>m. n \<le> Suc m \<longrightarrow> f m = 0 \<Longrightarrow> f NSsums (sum f {..<n})"
 | 
| 140 | by (auto simp add: sums_NSsums_iff [symmetric] not_le[symmetric] intro!: sums_finite) | |
| 27468 | 141 | |
| 142 | lemma NSsummable_NSCauchy: | |
| 75866 | 143 | "NSsummable f \<longleftrightarrow> (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. \<bar>sumhr (M, N, f)\<bar> \<approx> 0)" (is "?L=?R") | 
| 144 | proof - | |
| 145 | have "?L = (\<forall>M\<in>HNatInfinite. \<forall>N\<in>HNatInfinite. sumhr (0, M, f) \<approx> sumhr (0, N, f))" | |
| 146 | by (auto simp add: summable_iff_convergent convergent_NSconvergent_iff NSCauchy_def starfunNat_sumr | |
| 147 | simp flip: NSCauchy_NSconvergent_iff summable_NSsummable_iff atLeast0LessThan) | |
| 148 | also have "... \<longleftrightarrow> ?R" | |
| 149 | by (metis approx_hrabs_zero_cancel approx_minus_iff approx_refl approx_sym linorder_less_linear sumhr_hrabs_approx sumhr_split_diff) | |
| 150 | finally show ?thesis . | |
| 151 | qed | |
| 27468 | 152 | |
| 64604 | 153 | text \<open>Terms of a convergent series tend to zero.\<close> | 
| 154 | lemma NSsummable_NSLIMSEQ_zero: "NSsummable f \<Longrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S 0" | |
| 75866 | 155 | by (metis HNatInfinite_add NSLIMSEQ_def NSsummable_NSCauchy approx_hrabs_zero_cancel star_of_zero sumhr_Suc) | 
| 27468 | 156 | |
| 64604 | 157 | text \<open>Nonstandard comparison test.\<close> | 
| 158 | lemma NSsummable_comparison_test: "\<exists>N. \<forall>n. N \<le> n \<longrightarrow> \<bar>f n\<bar> \<le> g n \<Longrightarrow> NSsummable g \<Longrightarrow> NSsummable f" | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
64604diff
changeset | 159 | by (metis real_norm_def summable_NSsummable_iff summable_comparison_test) | 
| 27468 | 160 | |
| 161 | lemma NSsummable_rabs_comparison_test: | |
| 64604 | 162 | "\<exists>N. \<forall>n. N \<le> n \<longrightarrow> \<bar>f n\<bar> \<le> g n \<Longrightarrow> NSsummable g \<Longrightarrow> NSsummable (\<lambda>k. \<bar>f k\<bar>)" | 
| 163 | by (rule NSsummable_comparison_test) auto | |
| 27468 | 164 | |
| 165 | end |