| author | paulson <lp15@cam.ac.uk> | 
| Thu, 11 Apr 2019 22:37:49 +0100 | |
| changeset 70131 | c6e1a4806f49 | 
| parent 70113 | c8deb8ba6d05 | 
| child 70136 | f03a01a18c6e | 
| permissions | -rw-r--r-- | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 1 | (*File: HOL/Analysis/Infinite_Product.thy | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 2 | Author: Manuel Eberl & LC Paulson | 
| 66277 | 3 | |
| 4 | Basic results about convergence and absolute convergence of infinite products | |
| 5 | and their connection to summability. | |
| 6 | *) | |
| 7 | section \<open>Infinite Products\<close> | |
| 8 | theory Infinite_Products | |
| 68585 | 9 | imports Topology_Euclidean_Space Complex_Transcendental | 
| 66277 | 10 | begin | 
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 11 | |
| 68651 | 12 | subsection%unimportant \<open>Preliminaries\<close> | 
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 13 | |
| 66277 | 14 | lemma sum_le_prod: | 
| 15 | fixes f :: "'a \<Rightarrow> 'b :: linordered_semidom" | |
| 16 | assumes "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> 0" | |
| 17 | shows "sum f A \<le> (\<Prod>x\<in>A. 1 + f x)" | |
| 18 | using assms | |
| 19 | proof (induction A rule: infinite_finite_induct) | |
| 20 | case (insert x A) | |
| 21 | from insert.hyps have "sum f A + f x * (\<Prod>x\<in>A. 1) \<le> (\<Prod>x\<in>A. 1 + f x) + f x * (\<Prod>x\<in>A. 1 + f x)" | |
| 22 | by (intro add_mono insert mult_left_mono prod_mono) (auto intro: insert.prems) | |
| 23 | with insert.hyps show ?case by (simp add: algebra_simps) | |
| 24 | qed simp_all | |
| 25 | ||
| 26 | lemma prod_le_exp_sum: | |
| 27 | fixes f :: "'a \<Rightarrow> real" | |
| 28 | assumes "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> 0" | |
| 29 | shows "prod (\<lambda>x. 1 + f x) A \<le> exp (sum f A)" | |
| 30 | using assms | |
| 31 | proof (induction A rule: infinite_finite_induct) | |
| 32 | case (insert x A) | |
| 33 | have "(1 + f x) * (\<Prod>x\<in>A. 1 + f x) \<le> exp (f x) * exp (sum f A)" | |
| 34 | using insert.prems by (intro mult_mono insert prod_nonneg exp_ge_add_one_self) auto | |
| 35 | with insert.hyps show ?case by (simp add: algebra_simps exp_add) | |
| 36 | qed simp_all | |
| 37 | ||
| 38 | lemma lim_ln_1_plus_x_over_x_at_0: "(\<lambda>x::real. ln (1 + x) / x) \<midarrow>0\<rightarrow> 1" | |
| 39 | proof (rule lhopital) | |
| 40 | show "(\<lambda>x::real. ln (1 + x)) \<midarrow>0\<rightarrow> 0" | |
| 41 | by (rule tendsto_eq_intros refl | simp)+ | |
| 42 |   have "eventually (\<lambda>x::real. x \<in> {-1/2<..<1/2}) (nhds 0)"
 | |
| 43 | by (rule eventually_nhds_in_open) auto | |
| 44 |   hence *: "eventually (\<lambda>x::real. x \<in> {-1/2<..<1/2}) (at 0)"
 | |
| 45 | by (rule filter_leD [rotated]) (simp_all add: at_within_def) | |
| 46 | show "eventually (\<lambda>x::real. ((\<lambda>x. ln (1 + x)) has_field_derivative inverse (1 + x)) (at x)) (at 0)" | |
| 47 | using * by eventually_elim (auto intro!: derivative_eq_intros simp: field_simps) | |
| 48 | show "eventually (\<lambda>x::real. ((\<lambda>x. x) has_field_derivative 1) (at x)) (at 0)" | |
| 49 | using * by eventually_elim (auto intro!: derivative_eq_intros simp: field_simps) | |
| 50 | show "\<forall>\<^sub>F x in at 0. x \<noteq> 0" by (auto simp: at_within_def eventually_inf_principal) | |
| 51 | show "(\<lambda>x::real. inverse (1 + x) / 1) \<midarrow>0\<rightarrow> 1" | |
| 52 | by (rule tendsto_eq_intros refl | simp)+ | |
| 53 | qed auto | |
| 54 | ||
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 55 | subsection\<open>Definitions and basic properties\<close> | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 56 | |
| 68651 | 57 | definition%important raw_has_prod :: "[nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}, nat, 'a] \<Rightarrow> bool" 
 | 
| 68361 | 58 | where "raw_has_prod f M p \<equiv> (\<lambda>n. \<Prod>i\<le>n. f (i+M)) \<longlonglongrightarrow> p \<and> p \<noteq> 0" | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 59 | |
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 60 | text\<open>The nonzero and zero cases, as in \emph{Complex Analysis} by Joseph Bak and Donald J.Newman, page 241\<close>
 | 
| 69565 
1daf07b65385
retain important whitespace after 'text' that is suppressed, but swallows adjacent whitespace;
 wenzelm parents: 
69529diff
changeset | 61 | text%important \<open>%whitespace\<close> | 
| 68651 | 62 | definition%important | 
| 63 |   has_prod :: "(nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}) \<Rightarrow> 'a \<Rightarrow> bool" (infixr "has'_prod" 80)
 | |
| 68361 | 64 | where "f has_prod p \<equiv> raw_has_prod f 0 p \<or> (\<exists>i q. p = 0 \<and> f i = 0 \<and> raw_has_prod f (Suc i) q)" | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 65 | |
| 68651 | 66 | definition%important convergent_prod :: "(nat \<Rightarrow> 'a :: {t2_space,comm_semiring_1}) \<Rightarrow> bool" where
 | 
| 68361 | 67 | "convergent_prod f \<equiv> \<exists>M p. raw_has_prod f M p" | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 68 | |
| 68651 | 69 | definition%important prodinf :: "(nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}) \<Rightarrow> 'a"
 | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 70 | (binder "\<Prod>" 10) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 71 | where "prodinf f = (THE p. f has_prod p)" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 72 | |
| 68361 | 73 | lemmas prod_defs = raw_has_prod_def has_prod_def convergent_prod_def prodinf_def | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 74 | |
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 75 | lemma has_prod_subst[trans]: "f = g \<Longrightarrow> g has_prod z \<Longrightarrow> f has_prod z" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 76 | by simp | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 77 | |
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 78 | lemma has_prod_cong: "(\<And>n. f n = g n) \<Longrightarrow> f has_prod c \<longleftrightarrow> g has_prod c" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 79 | by presburger | 
| 66277 | 80 | |
| 68361 | 81 | lemma raw_has_prod_nonzero [simp]: "\<not> raw_has_prod f M 0" | 
| 82 | by (simp add: raw_has_prod_def) | |
| 68071 | 83 | |
| 68361 | 84 | lemma raw_has_prod_eq_0: | 
| 85 |   fixes f :: "nat \<Rightarrow> 'a::{semidom,t2_space}"
 | |
| 86 | assumes p: "raw_has_prod f m p" and i: "f i = 0" "i \<ge> m" | |
| 68136 | 87 | shows "p = 0" | 
| 88 | proof - | |
| 89 | have eq0: "(\<Prod>k\<le>n. f (k+m)) = 0" if "i - m \<le> n" for n | |
| 68361 | 90 | proof - | 
| 91 | have "\<exists>k\<le>n. f (k + m) = 0" | |
| 92 | using i that by auto | |
| 93 | then show ?thesis | |
| 94 | by auto | |
| 95 | qed | |
| 68136 | 96 | have "(\<lambda>n. \<Prod>i\<le>n. f (i + m)) \<longlonglongrightarrow> 0" | 
| 97 | by (rule LIMSEQ_offset [where k = "i-m"]) (simp add: eq0) | |
| 98 | with p show ?thesis | |
| 68361 | 99 | unfolding raw_has_prod_def | 
| 68136 | 100 | using LIMSEQ_unique by blast | 
| 101 | qed | |
| 102 | ||
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 103 | lemma raw_has_prod_Suc: | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 104 | "raw_has_prod f (Suc M) a \<longleftrightarrow> raw_has_prod (\<lambda>n. f (Suc n)) M a" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 105 | unfolding raw_has_prod_def by auto | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 106 | |
| 68361 | 107 | lemma has_prod_0_iff: "f has_prod 0 \<longleftrightarrow> (\<exists>i. f i = 0 \<and> (\<exists>p. raw_has_prod f (Suc i) p))" | 
| 68071 | 108 | by (simp add: has_prod_def) | 
| 68136 | 109 | |
| 110 | lemma has_prod_unique2: | |
| 68361 | 111 |   fixes f :: "nat \<Rightarrow> 'a::{semidom,t2_space}"
 | 
| 68136 | 112 | assumes "f has_prod a" "f has_prod b" shows "a = b" | 
| 113 | using assms | |
| 68361 | 114 | by (auto simp: has_prod_def raw_has_prod_eq_0) (meson raw_has_prod_def sequentially_bot tendsto_unique) | 
| 68136 | 115 | |
| 116 | lemma has_prod_unique: | |
| 68361 | 117 |   fixes f :: "nat \<Rightarrow> 'a :: {semidom,t2_space}"
 | 
| 68136 | 118 | shows "f has_prod s \<Longrightarrow> s = prodinf f" | 
| 119 | by (simp add: has_prod_unique2 prodinf_def the_equality) | |
| 68071 | 120 | |
| 66277 | 121 | lemma convergent_prod_altdef: | 
| 122 |   fixes f :: "nat \<Rightarrow> 'a :: {t2_space,comm_semiring_1}"
 | |
| 123 | shows "convergent_prod f \<longleftrightarrow> (\<exists>M L. (\<forall>n\<ge>M. f n \<noteq> 0) \<and> (\<lambda>n. \<Prod>i\<le>n. f (i+M)) \<longlonglongrightarrow> L \<and> L \<noteq> 0)" | |
| 124 | proof | |
| 125 | assume "convergent_prod f" | |
| 126 | then obtain M L where *: "(\<lambda>n. \<Prod>i\<le>n. f (i+M)) \<longlonglongrightarrow> L" "L \<noteq> 0" | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 127 | by (auto simp: prod_defs) | 
| 66277 | 128 | have "f i \<noteq> 0" if "i \<ge> M" for i | 
| 129 | proof | |
| 130 | assume "f i = 0" | |
| 131 | have **: "eventually (\<lambda>n. (\<Prod>i\<le>n. f (i+M)) = 0) sequentially" | |
| 132 | using eventually_ge_at_top[of "i - M"] | |
| 133 | proof eventually_elim | |
| 134 | case (elim n) | |
| 135 | with \<open>f i = 0\<close> and \<open>i \<ge> M\<close> show ?case | |
| 136 | by (auto intro!: bexI[of _ "i - M"] prod_zero) | |
| 137 | qed | |
| 138 | have "(\<lambda>n. (\<Prod>i\<le>n. f (i+M))) \<longlonglongrightarrow> 0" | |
| 139 | unfolding filterlim_iff | |
| 140 | by (auto dest!: eventually_nhds_x_imp_x intro!: eventually_mono[OF **]) | |
| 141 | from tendsto_unique[OF _ this *(1)] and *(2) | |
| 142 | show False by simp | |
| 143 | qed | |
| 144 | with * show "(\<exists>M L. (\<forall>n\<ge>M. f n \<noteq> 0) \<and> (\<lambda>n. \<Prod>i\<le>n. f (i+M)) \<longlonglongrightarrow> L \<and> L \<noteq> 0)" | |
| 145 | by blast | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 146 | qed (auto simp: prod_defs) | 
| 66277 | 147 | |
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 148 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 149 | subsection\<open>Absolutely convergent products\<close> | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 150 | |
| 68651 | 151 | definition%important abs_convergent_prod :: "(nat \<Rightarrow> _) \<Rightarrow> bool" where | 
| 66277 | 152 | "abs_convergent_prod f \<longleftrightarrow> convergent_prod (\<lambda>i. 1 + norm (f i - 1))" | 
| 153 | ||
| 154 | lemma abs_convergent_prodI: | |
| 155 | assumes "convergent (\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1))" | |
| 156 | shows "abs_convergent_prod f" | |
| 157 | proof - | |
| 158 | from assms obtain L where L: "(\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1)) \<longlonglongrightarrow> L" | |
| 159 | by (auto simp: convergent_def) | |
| 160 | have "L \<ge> 1" | |
| 161 | proof (rule tendsto_le) | |
| 162 | show "eventually (\<lambda>n. (\<Prod>i\<le>n. 1 + norm (f i - 1)) \<ge> 1) sequentially" | |
| 163 | proof (intro always_eventually allI) | |
| 164 | fix n | |
| 165 | have "(\<Prod>i\<le>n. 1 + norm (f i - 1)) \<ge> (\<Prod>i\<le>n. 1)" | |
| 166 | by (intro prod_mono) auto | |
| 167 | thus "(\<Prod>i\<le>n. 1 + norm (f i - 1)) \<ge> 1" by simp | |
| 168 | qed | |
| 169 | qed (use L in simp_all) | |
| 170 | hence "L \<noteq> 0" by auto | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 171 | with L show ?thesis unfolding abs_convergent_prod_def prod_defs | 
| 66277 | 172 | by (intro exI[of _ "0::nat"] exI[of _ L]) auto | 
| 173 | qed | |
| 174 | ||
| 175 | lemma | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 176 |   fixes f :: "nat \<Rightarrow> 'a :: {topological_semigroup_mult,t2_space,idom}"
 | 
| 66277 | 177 | assumes "convergent_prod f" | 
| 178 | shows convergent_prod_imp_convergent: "convergent (\<lambda>n. \<Prod>i\<le>n. f i)" | |
| 179 | and convergent_prod_to_zero_iff: "(\<lambda>n. \<Prod>i\<le>n. f i) \<longlonglongrightarrow> 0 \<longleftrightarrow> (\<exists>i. f i = 0)" | |
| 180 | proof - | |
| 181 | from assms obtain M L | |
| 182 | where M: "\<And>n. n \<ge> M \<Longrightarrow> f n \<noteq> 0" and "(\<lambda>n. \<Prod>i\<le>n. f (i + M)) \<longlonglongrightarrow> L" and "L \<noteq> 0" | |
| 183 | by (auto simp: convergent_prod_altdef) | |
| 184 | note this(2) | |
| 185 | also have "(\<lambda>n. \<Prod>i\<le>n. f (i + M)) = (\<lambda>n. \<Prod>i=M..M+n. f i)" | |
| 186 | by (intro ext prod.reindex_bij_witness[of _ "\<lambda>n. n - M" "\<lambda>n. n + M"]) auto | |
| 187 | finally have "(\<lambda>n. (\<Prod>i<M. f i) * (\<Prod>i=M..M+n. f i)) \<longlonglongrightarrow> (\<Prod>i<M. f i) * L" | |
| 188 | by (intro tendsto_mult tendsto_const) | |
| 189 |   also have "(\<lambda>n. (\<Prod>i<M. f i) * (\<Prod>i=M..M+n. f i)) = (\<lambda>n. (\<Prod>i\<in>{..<M}\<union>{M..M+n}. f i))"
 | |
| 190 | by (subst prod.union_disjoint) auto | |
| 191 |   also have "(\<lambda>n. {..<M} \<union> {M..M+n}) = (\<lambda>n. {..n+M})" by auto
 | |
| 192 |   finally have lim: "(\<lambda>n. prod f {..n}) \<longlonglongrightarrow> prod f {..<M} * L" 
 | |
| 193 | by (rule LIMSEQ_offset) | |
| 194 | thus "convergent (\<lambda>n. \<Prod>i\<le>n. f i)" | |
| 195 | by (auto simp: convergent_def) | |
| 196 | ||
| 197 | show "(\<lambda>n. \<Prod>i\<le>n. f i) \<longlonglongrightarrow> 0 \<longleftrightarrow> (\<exists>i. f i = 0)" | |
| 198 | proof | |
| 199 | assume "\<exists>i. f i = 0" | |
| 200 | then obtain i where "f i = 0" by auto | |
| 201 | moreover with M have "i < M" by (cases "i < M") auto | |
| 202 | ultimately have "(\<Prod>i<M. f i) = 0" by auto | |
| 203 | with lim show "(\<lambda>n. \<Prod>i\<le>n. f i) \<longlonglongrightarrow> 0" by simp | |
| 204 | next | |
| 205 | assume "(\<lambda>n. \<Prod>i\<le>n. f i) \<longlonglongrightarrow> 0" | |
| 206 | from tendsto_unique[OF _ this lim] and \<open>L \<noteq> 0\<close> | |
| 207 | show "\<exists>i. f i = 0" by auto | |
| 208 | qed | |
| 209 | qed | |
| 210 | ||
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 211 | lemma convergent_prod_iff_nz_lim: | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 212 |   fixes f :: "nat \<Rightarrow> 'a :: {topological_semigroup_mult,t2_space,idom}"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 213 | assumes "\<And>i. f i \<noteq> 0" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 214 | shows "convergent_prod f \<longleftrightarrow> (\<exists>L. (\<lambda>n. \<Prod>i\<le>n. f i) \<longlonglongrightarrow> L \<and> L \<noteq> 0)" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 215 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 216 | proof | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 217 | assume ?lhs then show ?rhs | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 218 | using assms convergentD convergent_prod_imp_convergent convergent_prod_to_zero_iff by blast | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 219 | next | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 220 | assume ?rhs then show ?lhs | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 221 | unfolding prod_defs | 
| 68138 | 222 | by (rule_tac x=0 in exI) auto | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 223 | qed | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 224 | |
| 68651 | 225 | lemma%important convergent_prod_iff_convergent: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 226 |   fixes f :: "nat \<Rightarrow> 'a :: {topological_semigroup_mult,t2_space,idom}"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 227 | assumes "\<And>i. f i \<noteq> 0" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 228 | shows "convergent_prod f \<longleftrightarrow> convergent (\<lambda>n. \<Prod>i\<le>n. f i) \<and> lim (\<lambda>n. \<Prod>i\<le>n. f i) \<noteq> 0" | 
| 68138 | 229 | by (force simp: convergent_prod_iff_nz_lim assms convergent_def limI) | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 230 | |
| 68527 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 231 | lemma bounded_imp_convergent_prod: | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 232 | fixes a :: "nat \<Rightarrow> real" | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 233 | assumes 1: "\<And>n. a n \<ge> 1" and bounded: "\<And>n. (\<Prod>i\<le>n. a i) \<le> B" | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 234 | shows "convergent_prod a" | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 235 | proof - | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 236 | have "bdd_above (range(\<lambda>n. \<Prod>i\<le>n. a i))" | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 237 | by (meson bdd_aboveI2 bounded) | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 238 | moreover have "incseq (\<lambda>n. \<Prod>i\<le>n. a i)" | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 239 | unfolding mono_def by (metis 1 prod_mono2 atMost_subset_iff dual_order.trans finite_atMost zero_le_one) | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 240 | ultimately obtain p where p: "(\<lambda>n. \<Prod>i\<le>n. a i) \<longlonglongrightarrow> p" | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 241 | using LIMSEQ_incseq_SUP by blast | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 242 | then have "p \<noteq> 0" | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 243 | by (metis "1" not_one_le_zero prod_ge_1 LIMSEQ_le_const) | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 244 | with 1 p show ?thesis | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 245 | by (metis convergent_prod_iff_nz_lim not_one_le_zero) | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 246 | qed | 
| 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 247 | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 248 | |
| 66277 | 249 | lemma abs_convergent_prod_altdef: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 250 |   fixes f :: "nat \<Rightarrow> 'a :: {one,real_normed_vector}"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 251 | shows "abs_convergent_prod f \<longleftrightarrow> convergent (\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1))" | 
| 66277 | 252 | proof | 
| 253 | assume "abs_convergent_prod f" | |
| 254 | thus "convergent (\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1))" | |
| 255 | by (auto simp: abs_convergent_prod_def intro!: convergent_prod_imp_convergent) | |
| 256 | qed (auto intro: abs_convergent_prodI) | |
| 257 | ||
| 69529 | 258 | lemma Weierstrass_prod_ineq: | 
| 66277 | 259 | fixes f :: "'a \<Rightarrow> real" | 
| 260 |   assumes "\<And>x. x \<in> A \<Longrightarrow> f x \<in> {0..1}"
 | |
| 261 | shows "1 - sum f A \<le> (\<Prod>x\<in>A. 1 - f x)" | |
| 262 | using assms | |
| 263 | proof (induction A rule: infinite_finite_induct) | |
| 264 | case (insert x A) | |
| 265 | from insert.hyps and insert.prems | |
| 266 | have "1 - sum f A + f x * (\<Prod>x\<in>A. 1 - f x) \<le> (\<Prod>x\<in>A. 1 - f x) + f x * (\<Prod>x\<in>A. 1)" | |
| 267 | by (intro insert.IH add_mono mult_left_mono prod_mono) auto | |
| 268 | with insert.hyps show ?case by (simp add: algebra_simps) | |
| 269 | qed simp_all | |
| 270 | ||
| 271 | lemma norm_prod_minus1_le_prod_minus1: | |
| 272 |   fixes f :: "nat \<Rightarrow> 'a :: {real_normed_div_algebra,comm_ring_1}"  
 | |
| 273 | shows "norm (prod (\<lambda>n. 1 + f n) A - 1) \<le> prod (\<lambda>n. 1 + norm (f n)) A - 1" | |
| 274 | proof (induction A rule: infinite_finite_induct) | |
| 275 | case (insert x A) | |
| 276 | from insert.hyps have | |
| 277 | "norm ((\<Prod>n\<in>insert x A. 1 + f n) - 1) = | |
| 278 | norm ((\<Prod>n\<in>A. 1 + f n) - 1 + f x * (\<Prod>n\<in>A. 1 + f n))" | |
| 279 | by (simp add: algebra_simps) | |
| 280 | also have "\<dots> \<le> norm ((\<Prod>n\<in>A. 1 + f n) - 1) + norm (f x * (\<Prod>n\<in>A. 1 + f n))" | |
| 281 | by (rule norm_triangle_ineq) | |
| 282 | also have "norm (f x * (\<Prod>n\<in>A. 1 + f n)) = norm (f x) * (\<Prod>x\<in>A. norm (1 + f x))" | |
| 283 | by (simp add: prod_norm norm_mult) | |
| 284 | also have "(\<Prod>x\<in>A. norm (1 + f x)) \<le> (\<Prod>x\<in>A. norm (1::'a) + norm (f x))" | |
| 285 | by (intro prod_mono norm_triangle_ineq ballI conjI) auto | |
| 286 | also have "norm (1::'a) = 1" by simp | |
| 287 | also note insert.IH | |
| 288 | also have "(\<Prod>n\<in>A. 1 + norm (f n)) - 1 + norm (f x) * (\<Prod>x\<in>A. 1 + norm (f x)) = | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 289 | (\<Prod>n\<in>insert x A. 1 + norm (f n)) - 1" | 
| 66277 | 290 | using insert.hyps by (simp add: algebra_simps) | 
| 291 | finally show ?case by - (simp_all add: mult_left_mono) | |
| 292 | qed simp_all | |
| 293 | ||
| 294 | lemma convergent_prod_imp_ev_nonzero: | |
| 295 |   fixes f :: "nat \<Rightarrow> 'a :: {t2_space,comm_semiring_1}"
 | |
| 296 | assumes "convergent_prod f" | |
| 297 | shows "eventually (\<lambda>n. f n \<noteq> 0) sequentially" | |
| 298 | using assms by (auto simp: eventually_at_top_linorder convergent_prod_altdef) | |
| 299 | ||
| 300 | lemma convergent_prod_imp_LIMSEQ: | |
| 301 |   fixes f :: "nat \<Rightarrow> 'a :: {real_normed_field}"
 | |
| 302 | assumes "convergent_prod f" | |
| 303 | shows "f \<longlonglongrightarrow> 1" | |
| 304 | proof - | |
| 305 | from assms obtain M L where L: "(\<lambda>n. \<Prod>i\<le>n. f (i+M)) \<longlonglongrightarrow> L" "\<And>n. n \<ge> M \<Longrightarrow> f n \<noteq> 0" "L \<noteq> 0" | |
| 306 | by (auto simp: convergent_prod_altdef) | |
| 307 | hence L': "(\<lambda>n. \<Prod>i\<le>Suc n. f (i+M)) \<longlonglongrightarrow> L" by (subst filterlim_sequentially_Suc) | |
| 308 | have "(\<lambda>n. (\<Prod>i\<le>Suc n. f (i+M)) / (\<Prod>i\<le>n. f (i+M))) \<longlonglongrightarrow> L / L" | |
| 309 | using L L' by (intro tendsto_divide) simp_all | |
| 310 | also from L have "L / L = 1" by simp | |
| 311 | also have "(\<lambda>n. (\<Prod>i\<le>Suc n. f (i+M)) / (\<Prod>i\<le>n. f (i+M))) = (\<lambda>n. f (n + Suc M))" | |
| 312 | using assms L by (auto simp: fun_eq_iff atMost_Suc) | |
| 313 | finally show ?thesis by (rule LIMSEQ_offset) | |
| 314 | qed | |
| 315 | ||
| 316 | lemma abs_convergent_prod_imp_summable: | |
| 317 | fixes f :: "nat \<Rightarrow> 'a :: real_normed_div_algebra" | |
| 318 | assumes "abs_convergent_prod f" | |
| 319 | shows "summable (\<lambda>i. norm (f i - 1))" | |
| 320 | proof - | |
| 321 | from assms have "convergent (\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1))" | |
| 322 | unfolding abs_convergent_prod_def by (rule convergent_prod_imp_convergent) | |
| 323 | then obtain L where L: "(\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1)) \<longlonglongrightarrow> L" | |
| 324 | unfolding convergent_def by blast | |
| 325 | have "convergent (\<lambda>n. \<Sum>i\<le>n. norm (f i - 1))" | |
| 326 | proof (rule Bseq_monoseq_convergent) | |
| 327 | have "eventually (\<lambda>n. (\<Prod>i\<le>n. 1 + norm (f i - 1)) < L + 1) sequentially" | |
| 328 | using L(1) by (rule order_tendstoD) simp_all | |
| 329 | hence "\<forall>\<^sub>F x in sequentially. norm (\<Sum>i\<le>x. norm (f i - 1)) \<le> L + 1" | |
| 330 | proof eventually_elim | |
| 331 | case (elim n) | |
| 332 | have "norm (\<Sum>i\<le>n. norm (f i - 1)) = (\<Sum>i\<le>n. norm (f i - 1))" | |
| 333 | unfolding real_norm_def by (intro abs_of_nonneg sum_nonneg) simp_all | |
| 334 | also have "\<dots> \<le> (\<Prod>i\<le>n. 1 + norm (f i - 1))" by (rule sum_le_prod) auto | |
| 335 | also have "\<dots> < L + 1" by (rule elim) | |
| 336 | finally show ?case by simp | |
| 337 | qed | |
| 338 | thus "Bseq (\<lambda>n. \<Sum>i\<le>n. norm (f i - 1))" by (rule BfunI) | |
| 339 | next | |
| 340 | show "monoseq (\<lambda>n. \<Sum>i\<le>n. norm (f i - 1))" | |
| 341 | by (rule mono_SucI1) auto | |
| 342 | qed | |
| 343 | thus "summable (\<lambda>i. norm (f i - 1))" by (simp add: summable_iff_convergent') | |
| 344 | qed | |
| 345 | ||
| 346 | lemma summable_imp_abs_convergent_prod: | |
| 347 | fixes f :: "nat \<Rightarrow> 'a :: real_normed_div_algebra" | |
| 348 | assumes "summable (\<lambda>i. norm (f i - 1))" | |
| 349 | shows "abs_convergent_prod f" | |
| 350 | proof (intro abs_convergent_prodI Bseq_monoseq_convergent) | |
| 351 | show "monoseq (\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1))" | |
| 352 | by (intro mono_SucI1) | |
| 353 | (auto simp: atMost_Suc algebra_simps intro!: mult_nonneg_nonneg prod_nonneg) | |
| 354 | next | |
| 355 | show "Bseq (\<lambda>n. \<Prod>i\<le>n. 1 + norm (f i - 1))" | |
| 356 | proof (rule Bseq_eventually_mono) | |
| 357 | show "eventually (\<lambda>n. norm (\<Prod>i\<le>n. 1 + norm (f i - 1)) \<le> | |
| 358 | norm (exp (\<Sum>i\<le>n. norm (f i - 1)))) sequentially" | |
| 359 | by (intro always_eventually allI) (auto simp: abs_prod exp_sum intro!: prod_mono) | |
| 360 | next | |
| 361 | from assms have "(\<lambda>n. \<Sum>i\<le>n. norm (f i - 1)) \<longlonglongrightarrow> (\<Sum>i. norm (f i - 1))" | |
| 362 | using sums_def_le by blast | |
| 363 | hence "(\<lambda>n. exp (\<Sum>i\<le>n. norm (f i - 1))) \<longlonglongrightarrow> exp (\<Sum>i. norm (f i - 1))" | |
| 364 | by (rule tendsto_exp) | |
| 365 | hence "convergent (\<lambda>n. exp (\<Sum>i\<le>n. norm (f i - 1)))" | |
| 366 | by (rule convergentI) | |
| 367 | thus "Bseq (\<lambda>n. exp (\<Sum>i\<le>n. norm (f i - 1)))" | |
| 368 | by (rule convergent_imp_Bseq) | |
| 369 | qed | |
| 370 | qed | |
| 371 | ||
| 68651 | 372 | theorem abs_convergent_prod_conv_summable: | 
| 66277 | 373 | fixes f :: "nat \<Rightarrow> 'a :: real_normed_div_algebra" | 
| 374 | shows "abs_convergent_prod f \<longleftrightarrow> summable (\<lambda>i. norm (f i - 1))" | |
| 375 | by (blast intro: abs_convergent_prod_imp_summable summable_imp_abs_convergent_prod) | |
| 376 | ||
| 377 | lemma abs_convergent_prod_imp_LIMSEQ: | |
| 378 |   fixes f :: "nat \<Rightarrow> 'a :: {comm_ring_1,real_normed_div_algebra}"
 | |
| 379 | assumes "abs_convergent_prod f" | |
| 380 | shows "f \<longlonglongrightarrow> 1" | |
| 381 | proof - | |
| 382 | from assms have "summable (\<lambda>n. norm (f n - 1))" | |
| 383 | by (rule abs_convergent_prod_imp_summable) | |
| 384 | from summable_LIMSEQ_zero[OF this] have "(\<lambda>n. f n - 1) \<longlonglongrightarrow> 0" | |
| 385 | by (simp add: tendsto_norm_zero_iff) | |
| 386 | from tendsto_add[OF this tendsto_const[of 1]] show ?thesis by simp | |
| 387 | qed | |
| 388 | ||
| 389 | lemma abs_convergent_prod_imp_ev_nonzero: | |
| 390 |   fixes f :: "nat \<Rightarrow> 'a :: {comm_ring_1,real_normed_div_algebra}"
 | |
| 391 | assumes "abs_convergent_prod f" | |
| 392 | shows "eventually (\<lambda>n. f n \<noteq> 0) sequentially" | |
| 393 | proof - | |
| 394 | from assms have "f \<longlonglongrightarrow> 1" | |
| 395 | by (rule abs_convergent_prod_imp_LIMSEQ) | |
| 396 | hence "eventually (\<lambda>n. dist (f n) 1 < 1) at_top" | |
| 397 | by (auto simp: tendsto_iff) | |
| 398 | thus ?thesis by eventually_elim auto | |
| 399 | qed | |
| 400 | ||
| 68651 | 401 | subsection%unimportant \<open>Ignoring initial segments\<close> | 
| 402 | ||
| 66277 | 403 | lemma convergent_prod_offset: | 
| 404 | assumes "convergent_prod (\<lambda>n. f (n + m))" | |
| 405 | shows "convergent_prod f" | |
| 406 | proof - | |
| 407 | from assms obtain M L where "(\<lambda>n. \<Prod>k\<le>n. f (k + (M + m))) \<longlonglongrightarrow> L" "L \<noteq> 0" | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 408 | by (auto simp: prod_defs add.assoc) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 409 | thus "convergent_prod f" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 410 | unfolding prod_defs by blast | 
| 66277 | 411 | qed | 
| 412 | ||
| 413 | lemma abs_convergent_prod_offset: | |
| 414 | assumes "abs_convergent_prod (\<lambda>n. f (n + m))" | |
| 415 | shows "abs_convergent_prod f" | |
| 416 | using assms unfolding abs_convergent_prod_def by (rule convergent_prod_offset) | |
| 417 | ||
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 418 | |
| 68361 | 419 | lemma raw_has_prod_ignore_initial_segment: | 
| 420 | fixes f :: "nat \<Rightarrow> 'a :: real_normed_field" | |
| 421 | assumes "raw_has_prod f M p" "N \<ge> M" | |
| 422 | obtains q where "raw_has_prod f N q" | |
| 66277 | 423 | proof - | 
| 68361 | 424 | have p: "(\<lambda>n. \<Prod>k\<le>n. f (k + M)) \<longlonglongrightarrow> p" and "p \<noteq> 0" | 
| 425 | using assms by (auto simp: raw_has_prod_def) | |
| 426 | then have nz: "\<And>n. n \<ge> M \<Longrightarrow> f n \<noteq> 0" | |
| 427 | using assms by (auto simp: raw_has_prod_eq_0) | |
| 428 | define C where "C = (\<Prod>k<N-M. f (k + M))" | |
| 66277 | 429 | from nz have [simp]: "C \<noteq> 0" | 
| 430 | by (auto simp: C_def) | |
| 431 | ||
| 68361 | 432 | from p have "(\<lambda>i. \<Prod>k\<le>i + (N-M). f (k + M)) \<longlonglongrightarrow> p" | 
| 66277 | 433 | by (rule LIMSEQ_ignore_initial_segment) | 
| 68361 | 434 | also have "(\<lambda>i. \<Prod>k\<le>i + (N-M). f (k + M)) = (\<lambda>n. C * (\<Prod>k\<le>n. f (k + N)))" | 
| 66277 | 435 | proof (rule ext, goal_cases) | 
| 436 | case (1 n) | |
| 68361 | 437 |     have "{..n+(N-M)} = {..<(N-M)} \<union> {(N-M)..n+(N-M)}" by auto
 | 
| 438 | also have "(\<Prod>k\<in>\<dots>. f (k + M)) = C * (\<Prod>k=(N-M)..n+(N-M). f (k + M))" | |
| 66277 | 439 | unfolding C_def by (rule prod.union_disjoint) auto | 
| 68361 | 440 | also have "(\<Prod>k=(N-M)..n+(N-M). f (k + M)) = (\<Prod>k\<le>n. f (k + (N-M) + M))" | 
| 441 | by (intro ext prod.reindex_bij_witness[of _ "\<lambda>k. k + (N-M)" "\<lambda>k. k - (N-M)"]) auto | |
| 442 | finally show ?case | |
| 443 | using \<open>N \<ge> M\<close> by (simp add: add_ac) | |
| 66277 | 444 | qed | 
| 68361 | 445 | finally have "(\<lambda>n. C * (\<Prod>k\<le>n. f (k + N)) / C) \<longlonglongrightarrow> p / C" | 
| 66277 | 446 | by (intro tendsto_divide tendsto_const) auto | 
| 68361 | 447 | hence "(\<lambda>n. \<Prod>k\<le>n. f (k + N)) \<longlonglongrightarrow> p / C" by simp | 
| 448 | moreover from \<open>p \<noteq> 0\<close> have "p / C \<noteq> 0" by simp | |
| 449 | ultimately show ?thesis | |
| 450 | using raw_has_prod_def that by blast | |
| 66277 | 451 | qed | 
| 452 | ||
| 68651 | 453 | corollary%unimportant convergent_prod_ignore_initial_segment: | 
| 68361 | 454 | fixes f :: "nat \<Rightarrow> 'a :: real_normed_field" | 
| 455 | assumes "convergent_prod f" | |
| 456 | shows "convergent_prod (\<lambda>n. f (n + m))" | |
| 457 | using assms | |
| 458 | unfolding convergent_prod_def | |
| 459 | apply clarify | |
| 460 | apply (erule_tac N="M+m" in raw_has_prod_ignore_initial_segment) | |
| 461 | apply (auto simp add: raw_has_prod_def add_ac) | |
| 462 | done | |
| 463 | ||
| 68651 | 464 | corollary%unimportant convergent_prod_ignore_nonzero_segment: | 
| 68136 | 465 | fixes f :: "nat \<Rightarrow> 'a :: real_normed_field" | 
| 466 | assumes f: "convergent_prod f" and nz: "\<And>i. i \<ge> M \<Longrightarrow> f i \<noteq> 0" | |
| 68361 | 467 | shows "\<exists>p. raw_has_prod f M p" | 
| 68136 | 468 | using convergent_prod_ignore_initial_segment [OF f] | 
| 469 | by (metis convergent_LIMSEQ_iff convergent_prod_iff_convergent le_add_same_cancel2 nz prod_defs(1) zero_order(1)) | |
| 470 | ||
| 68651 | 471 | corollary%unimportant abs_convergent_prod_ignore_initial_segment: | 
| 66277 | 472 | assumes "abs_convergent_prod f" | 
| 473 | shows "abs_convergent_prod (\<lambda>n. f (n + m))" | |
| 474 | using assms unfolding abs_convergent_prod_def | |
| 475 | by (rule convergent_prod_ignore_initial_segment) | |
| 476 | ||
| 68651 | 477 | subsection\<open>More elementary properties\<close> | 
| 478 | ||
| 479 | theorem abs_convergent_prod_imp_convergent_prod: | |
| 66277 | 480 |   fixes f :: "nat \<Rightarrow> 'a :: {real_normed_div_algebra,complete_space,comm_ring_1}"
 | 
| 481 | assumes "abs_convergent_prod f" | |
| 482 | shows "convergent_prod f" | |
| 483 | proof - | |
| 484 | from assms have "eventually (\<lambda>n. f n \<noteq> 0) sequentially" | |
| 485 | by (rule abs_convergent_prod_imp_ev_nonzero) | |
| 486 | then obtain N where N: "f n \<noteq> 0" if "n \<ge> N" for n | |
| 487 | by (auto simp: eventually_at_top_linorder) | |
| 488 | let ?P = "\<lambda>n. \<Prod>i\<le>n. f (i + N)" and ?Q = "\<lambda>n. \<Prod>i\<le>n. 1 + norm (f (i + N) - 1)" | |
| 489 | ||
| 490 | have "Cauchy ?P" | |
| 491 | proof (rule CauchyI', goal_cases) | |
| 492 | case (1 \<epsilon>) | |
| 493 | from assms have "abs_convergent_prod (\<lambda>n. f (n + N))" | |
| 494 | by (rule abs_convergent_prod_ignore_initial_segment) | |
| 495 | hence "Cauchy ?Q" | |
| 496 | unfolding abs_convergent_prod_def | |
| 497 | by (intro convergent_Cauchy convergent_prod_imp_convergent) | |
| 498 | from CauchyD[OF this 1] obtain M where M: "norm (?Q m - ?Q n) < \<epsilon>" if "m \<ge> M" "n \<ge> M" for m n | |
| 499 | by blast | |
| 500 | show ?case | |
| 501 | proof (rule exI[of _ M], safe, goal_cases) | |
| 502 | case (1 m n) | |
| 503 | have "dist (?P m) (?P n) = norm (?P n - ?P m)" | |
| 504 | by (simp add: dist_norm norm_minus_commute) | |
| 505 |       also from 1 have "{..n} = {..m} \<union> {m<..n}" by auto
 | |
| 506 |       hence "norm (?P n - ?P m) = norm (?P m * (\<Prod>k\<in>{m<..n}. f (k + N)) - ?P m)"
 | |
| 507 | by (subst prod.union_disjoint [symmetric]) (auto simp: algebra_simps) | |
| 508 |       also have "\<dots> = norm (?P m * ((\<Prod>k\<in>{m<..n}. f (k + N)) - 1))"
 | |
| 509 | by (simp add: algebra_simps) | |
| 510 |       also have "\<dots> = (\<Prod>k\<le>m. norm (f (k + N))) * norm ((\<Prod>k\<in>{m<..n}. f (k + N)) - 1)"
 | |
| 511 | by (simp add: norm_mult prod_norm) | |
| 512 |       also have "\<dots> \<le> ?Q m * ((\<Prod>k\<in>{m<..n}. 1 + norm (f (k + N) - 1)) - 1)"
 | |
| 513 |         using norm_prod_minus1_le_prod_minus1[of "\<lambda>k. f (k + N) - 1" "{m<..n}"]
 | |
| 514 | norm_triangle_ineq[of 1 "f k - 1" for k] | |
| 515 | by (intro mult_mono prod_mono ballI conjI norm_prod_minus1_le_prod_minus1 prod_nonneg) auto | |
| 516 |       also have "\<dots> = ?Q m * (\<Prod>k\<in>{m<..n}. 1 + norm (f (k + N) - 1)) - ?Q m"
 | |
| 517 | by (simp add: algebra_simps) | |
| 518 |       also have "?Q m * (\<Prod>k\<in>{m<..n}. 1 + norm (f (k + N) - 1)) = 
 | |
| 519 |                    (\<Prod>k\<in>{..m}\<union>{m<..n}. 1 + norm (f (k + N) - 1))"
 | |
| 520 | by (rule prod.union_disjoint [symmetric]) auto | |
| 521 |       also from 1 have "{..m}\<union>{m<..n} = {..n}" by auto
 | |
| 522 | also have "?Q n - ?Q m \<le> norm (?Q n - ?Q m)" by simp | |
| 523 | also from 1 have "\<dots> < \<epsilon>" by (intro M) auto | |
| 524 | finally show ?case . | |
| 525 | qed | |
| 526 | qed | |
| 527 | hence conv: "convergent ?P" by (rule Cauchy_convergent) | |
| 528 | then obtain L where L: "?P \<longlonglongrightarrow> L" | |
| 529 | by (auto simp: convergent_def) | |
| 530 | ||
| 531 | have "L \<noteq> 0" | |
| 532 | proof | |
| 533 | assume [simp]: "L = 0" | |
| 534 | from tendsto_norm[OF L] have limit: "(\<lambda>n. \<Prod>k\<le>n. norm (f (k + N))) \<longlonglongrightarrow> 0" | |
| 535 | by (simp add: prod_norm) | |
| 536 | ||
| 537 | from assms have "(\<lambda>n. f (n + N)) \<longlonglongrightarrow> 1" | |
| 538 | by (intro abs_convergent_prod_imp_LIMSEQ abs_convergent_prod_ignore_initial_segment) | |
| 539 | hence "eventually (\<lambda>n. norm (f (n + N) - 1) < 1) sequentially" | |
| 540 | by (auto simp: tendsto_iff dist_norm) | |
| 541 | then obtain M0 where M0: "norm (f (n + N) - 1) < 1" if "n \<ge> M0" for n | |
| 542 | by (auto simp: eventually_at_top_linorder) | |
| 543 | ||
| 544 |     {
 | |
| 545 | fix M assume M: "M \<ge> M0" | |
| 546 | with M0 have M: "norm (f (n + N) - 1) < 1" if "n \<ge> M" for n using that by simp | |
| 547 | ||
| 548 | have "(\<lambda>n. \<Prod>k\<le>n. 1 - norm (f (k+M+N) - 1)) \<longlonglongrightarrow> 0" | |
| 549 | proof (rule tendsto_sandwich) | |
| 550 | show "eventually (\<lambda>n. (\<Prod>k\<le>n. 1 - norm (f (k+M+N) - 1)) \<ge> 0) sequentially" | |
| 551 | using M by (intro always_eventually prod_nonneg allI ballI) (auto intro: less_imp_le) | |
| 552 | have "norm (1::'a) - norm (f (i + M + N) - 1) \<le> norm (f (i + M + N))" for i | |
| 553 | using norm_triangle_ineq3[of "f (i + M + N)" 1] by simp | |
| 554 | thus "eventually (\<lambda>n. (\<Prod>k\<le>n. 1 - norm (f (k+M+N) - 1)) \<le> (\<Prod>k\<le>n. norm (f (k+M+N)))) at_top" | |
| 555 | using M by (intro always_eventually allI prod_mono ballI conjI) (auto intro: less_imp_le) | |
| 556 | ||
| 557 | define C where "C = (\<Prod>k<M. norm (f (k + N)))" | |
| 558 | from N have [simp]: "C \<noteq> 0" by (auto simp: C_def) | |
| 559 | from L have "(\<lambda>n. norm (\<Prod>k\<le>n+M. f (k + N))) \<longlonglongrightarrow> 0" | |
| 560 | by (intro LIMSEQ_ignore_initial_segment) (simp add: tendsto_norm_zero_iff) | |
| 561 | also have "(\<lambda>n. norm (\<Prod>k\<le>n+M. f (k + N))) = (\<lambda>n. C * (\<Prod>k\<le>n. norm (f (k + M + N))))" | |
| 562 | proof (rule ext, goal_cases) | |
| 563 | case (1 n) | |
| 564 |           have "{..n+M} = {..<M} \<union> {M..n+M}" by auto
 | |
| 565 | also have "norm (\<Prod>k\<in>\<dots>. f (k + N)) = C * norm (\<Prod>k=M..n+M. f (k + N))" | |
| 566 | unfolding C_def by (subst prod.union_disjoint) (auto simp: norm_mult prod_norm) | |
| 567 | also have "(\<Prod>k=M..n+M. f (k + N)) = (\<Prod>k\<le>n. f (k + N + M))" | |
| 568 | by (intro prod.reindex_bij_witness[of _ "\<lambda>i. i + M" "\<lambda>i. i - M"]) auto | |
| 569 | finally show ?case by (simp add: add_ac prod_norm) | |
| 570 | qed | |
| 571 | finally have "(\<lambda>n. C * (\<Prod>k\<le>n. norm (f (k + M + N))) / C) \<longlonglongrightarrow> 0 / C" | |
| 572 | by (intro tendsto_divide tendsto_const) auto | |
| 573 | thus "(\<lambda>n. \<Prod>k\<le>n. norm (f (k + M + N))) \<longlonglongrightarrow> 0" by simp | |
| 574 | qed simp_all | |
| 575 | ||
| 576 | have "1 - (\<Sum>i. norm (f (i + M + N) - 1)) \<le> 0" | |
| 577 | proof (rule tendsto_le) | |
| 578 | show "eventually (\<lambda>n. 1 - (\<Sum>k\<le>n. norm (f (k+M+N) - 1)) \<le> | |
| 579 | (\<Prod>k\<le>n. 1 - norm (f (k+M+N) - 1))) at_top" | |
| 69529 | 580 | using M by (intro always_eventually allI Weierstrass_prod_ineq) (auto intro: less_imp_le) | 
| 66277 | 581 | show "(\<lambda>n. \<Prod>k\<le>n. 1 - norm (f (k+M+N) - 1)) \<longlonglongrightarrow> 0" by fact | 
| 582 | show "(\<lambda>n. 1 - (\<Sum>k\<le>n. norm (f (k + M + N) - 1))) | |
| 583 | \<longlonglongrightarrow> 1 - (\<Sum>i. norm (f (i + M + N) - 1))" | |
| 584 | by (intro tendsto_intros summable_LIMSEQ' summable_ignore_initial_segment | |
| 585 | abs_convergent_prod_imp_summable assms) | |
| 586 | qed simp_all | |
| 587 | hence "(\<Sum>i. norm (f (i + M + N) - 1)) \<ge> 1" by simp | |
| 588 | also have "\<dots> + (\<Sum>i<M. norm (f (i + N) - 1)) = (\<Sum>i. norm (f (i + N) - 1))" | |
| 589 | by (intro suminf_split_initial_segment [symmetric] summable_ignore_initial_segment | |
| 590 | abs_convergent_prod_imp_summable assms) | |
| 591 | finally have "1 + (\<Sum>i<M. norm (f (i + N) - 1)) \<le> (\<Sum>i. norm (f (i + N) - 1))" by simp | |
| 592 | } note * = this | |
| 593 | ||
| 594 | have "1 + (\<Sum>i. norm (f (i + N) - 1)) \<le> (\<Sum>i. norm (f (i + N) - 1))" | |
| 595 | proof (rule tendsto_le) | |
| 596 | show "(\<lambda>M. 1 + (\<Sum>i<M. norm (f (i + N) - 1))) \<longlonglongrightarrow> 1 + (\<Sum>i. norm (f (i + N) - 1))" | |
| 597 | by (intro tendsto_intros summable_LIMSEQ summable_ignore_initial_segment | |
| 598 | abs_convergent_prod_imp_summable assms) | |
| 599 | show "eventually (\<lambda>M. 1 + (\<Sum>i<M. norm (f (i + N) - 1)) \<le> (\<Sum>i. norm (f (i + N) - 1))) at_top" | |
| 600 | using eventually_ge_at_top[of M0] by eventually_elim (use * in auto) | |
| 601 | qed simp_all | |
| 602 | thus False by simp | |
| 603 | qed | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 604 | with L show ?thesis by (auto simp: prod_defs) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 605 | qed | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 606 | |
| 68361 | 607 | lemma raw_has_prod_cases: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 608 |   fixes f :: "nat \<Rightarrow> 'a :: {idom,topological_semigroup_mult,t2_space}"
 | 
| 68361 | 609 | assumes "raw_has_prod f M p" | 
| 610 | obtains i where "i<M" "f i = 0" | p where "raw_has_prod f 0 p" | |
| 68136 | 611 | proof - | 
| 612 | have "(\<lambda>n. \<Prod>i\<le>n. f (i + M)) \<longlonglongrightarrow> p" "p \<noteq> 0" | |
| 68361 | 613 | using assms unfolding raw_has_prod_def by blast+ | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 614 |   then have "(\<lambda>n. prod f {..<M} * (\<Prod>i\<le>n. f (i + M))) \<longlonglongrightarrow> prod f {..<M} * p"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 615 | by (metis tendsto_mult_left) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 616 |   moreover have "prod f {..<M} * (\<Prod>i\<le>n. f (i + M)) = prod f {..n+M}" for n
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 617 | proof - | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 618 |     have "{..n+M} = {..<M} \<union> {M..n+M}"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 619 | by auto | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 620 |     then have "prod f {..n+M} = prod f {..<M} * prod f {M..n+M}"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 621 | by simp (subst prod.union_disjoint; force) | 
| 68138 | 622 |     also have "\<dots> = prod f {..<M} * (\<Prod>i\<le>n. f (i + M))"
 | 
| 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: 
69565diff
changeset | 623 | by (metis (mono_tags, lifting) add.left_neutral atMost_atLeast0 prod.shift_bounds_cl_nat_ivl) | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 624 | finally show ?thesis by metis | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 625 | qed | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 626 |   ultimately have "(\<lambda>n. prod f {..n}) \<longlonglongrightarrow> prod f {..<M} * p"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 627 | by (auto intro: LIMSEQ_offset [where k=M]) | 
| 68361 | 628 |   then have "raw_has_prod f 0 (prod f {..<M} * p)" if "\<forall>i<M. f i \<noteq> 0"
 | 
| 629 | using \<open>p \<noteq> 0\<close> assms that by (auto simp: raw_has_prod_def) | |
| 68136 | 630 | then show thesis | 
| 631 | using that by blast | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 632 | qed | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 633 | |
| 68136 | 634 | corollary convergent_prod_offset_0: | 
| 635 |   fixes f :: "nat \<Rightarrow> 'a :: {idom,topological_semigroup_mult,t2_space}"
 | |
| 636 | assumes "convergent_prod f" "\<And>i. f i \<noteq> 0" | |
| 68361 | 637 | shows "\<exists>p. raw_has_prod f 0 p" | 
| 638 | using assms convergent_prod_def raw_has_prod_cases by blast | |
| 68136 | 639 | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 640 | lemma prodinf_eq_lim: | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 641 |   fixes f :: "nat \<Rightarrow> 'a :: {idom,topological_semigroup_mult,t2_space}"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 642 | assumes "convergent_prod f" "\<And>i. f i \<noteq> 0" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 643 | shows "prodinf f = lim (\<lambda>n. \<Prod>i\<le>n. f i)" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 644 | using assms convergent_prod_offset_0 [OF assms] | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 645 | by (simp add: prod_defs lim_def) (metis (no_types) assms(1) convergent_prod_to_zero_iff) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 646 | |
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 647 | lemma has_prod_one[simp, intro]: "(\<lambda>n. 1) has_prod 1" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 648 | unfolding prod_defs by auto | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 649 | |
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 650 | lemma convergent_prod_one[simp, intro]: "convergent_prod (\<lambda>n. 1)" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 651 | unfolding prod_defs by auto | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 652 | |
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 653 | lemma prodinf_cong: "(\<And>n. f n = g n) \<Longrightarrow> prodinf f = prodinf g" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 654 | by presburger | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 655 | |
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 656 | lemma convergent_prod_cong: | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 657 |   fixes f g :: "nat \<Rightarrow> 'a::{field,topological_semigroup_mult,t2_space}"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 658 | assumes ev: "eventually (\<lambda>x. f x = g x) sequentially" and f: "\<And>i. f i \<noteq> 0" and g: "\<And>i. g i \<noteq> 0" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 659 | shows "convergent_prod f = convergent_prod g" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 660 | proof - | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 661 | from assms obtain N where N: "\<forall>n\<ge>N. f n = g n" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 662 | by (auto simp: eventually_at_top_linorder) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 663 | define C where "C = (\<Prod>k<N. f k / g k)" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 664 | with g have "C \<noteq> 0" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 665 | by (simp add: f) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 666 |   have *: "eventually (\<lambda>n. prod f {..n} = C * prod g {..n}) sequentially"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 667 | using eventually_ge_at_top[of N] | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 668 | proof eventually_elim | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 669 | case (elim n) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 670 |     then have "{..n} = {..<N} \<union> {N..n}"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 671 | by auto | 
| 68138 | 672 |     also have "prod f \<dots> = prod f {..<N} * prod f {N..n}"
 | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 673 | by (intro prod.union_disjoint) auto | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 674 |     also from N have "prod f {N..n} = prod g {N..n}"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 675 | by (intro prod.cong) simp_all | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 676 |     also have "prod f {..<N} * prod g {N..n} = C * (prod g {..<N} * prod g {N..n})"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 677 | unfolding C_def by (simp add: g prod_dividef) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 678 |     also have "prod g {..<N} * prod g {N..n} = prod g ({..<N} \<union> {N..n})"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 679 | by (intro prod.union_disjoint [symmetric]) auto | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 680 |     also from elim have "{..<N} \<union> {N..n} = {..n}"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 681 | by auto | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 682 |     finally show "prod f {..n} = C * prod g {..n}" .
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 683 | qed | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 684 |   then have cong: "convergent (\<lambda>n. prod f {..n}) = convergent (\<lambda>n. C * prod g {..n})"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 685 | by (rule convergent_cong) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 686 | show ?thesis | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 687 | proof | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 688 | assume cf: "convergent_prod f" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 689 |     then have "\<not> (\<lambda>n. prod g {..n}) \<longlonglongrightarrow> 0"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 690 | using tendsto_mult_left * convergent_prod_to_zero_iff f filterlim_cong by fastforce | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 691 | then show "convergent_prod g" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 692 | by (metis convergent_mult_const_iff \<open>C \<noteq> 0\<close> cong cf convergent_LIMSEQ_iff convergent_prod_iff_convergent convergent_prod_imp_convergent g) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 693 | next | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 694 | assume cg: "convergent_prod g" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 695 |     have "\<exists>a. C * a \<noteq> 0 \<and> (\<lambda>n. prod g {..n}) \<longlonglongrightarrow> a"
 | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 696 | by (metis (no_types) \<open>C \<noteq> 0\<close> cg convergent_prod_iff_nz_lim divide_eq_0_iff g nonzero_mult_div_cancel_right) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 697 | then show "convergent_prod f" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 698 | using "*" tendsto_mult_left filterlim_cong | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 699 | by (fastforce simp add: convergent_prod_iff_nz_lim f) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
66277diff
changeset | 700 | qed | 
| 66277 | 701 | qed | 
| 702 | ||
| 68071 | 703 | lemma has_prod_finite: | 
| 68361 | 704 |   fixes f :: "nat \<Rightarrow> 'a::{semidom,t2_space}"
 | 
| 68071 | 705 | assumes [simp]: "finite N" | 
| 706 | and f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 1" | |
| 707 | shows "f has_prod (\<Prod>n\<in>N. f n)" | |
| 708 | proof - | |
| 709 |   have eq: "prod f {..n + Suc (Max N)} = prod f N" for n
 | |
| 710 | proof (rule prod.mono_neutral_right) | |
| 711 |     show "N \<subseteq> {..n + Suc (Max N)}"
 | |
| 68138 | 712 | by (auto simp: le_Suc_eq trans_le_add2) | 
| 68071 | 713 |     show "\<forall>i\<in>{..n + Suc (Max N)} - N. f i = 1"
 | 
| 714 | using f by blast | |
| 715 | qed auto | |
| 716 | show ?thesis | |
| 717 | proof (cases "\<forall>n\<in>N. f n \<noteq> 0") | |
| 718 | case True | |
| 719 | then have "prod f N \<noteq> 0" | |
| 720 | by simp | |
| 721 |     moreover have "(\<lambda>n. prod f {..n}) \<longlonglongrightarrow> prod f N"
 | |
| 722 | by (rule LIMSEQ_offset[of _ "Suc (Max N)"]) (simp add: eq atLeast0LessThan del: add_Suc_right) | |
| 723 | ultimately show ?thesis | |
| 68361 | 724 | by (simp add: raw_has_prod_def has_prod_def) | 
| 68071 | 725 | next | 
| 726 | case False | |
| 727 | then obtain k where "k \<in> N" "f k = 0" | |
| 728 | by auto | |
| 729 |     let ?Z = "{n \<in> N. f n = 0}"
 | |
| 730 | have maxge: "Max ?Z \<ge> n" if "f n = 0" for n | |
| 731 | using Max_ge [of ?Z] \<open>finite N\<close> \<open>f n = 0\<close> | |
| 732 | by (metis (mono_tags) Collect_mem_eq f finite_Collect_conjI mem_Collect_eq zero_neq_one) | |
| 733 |     let ?q = "prod f {Suc (Max ?Z)..Max N}"
 | |
| 734 | have [simp]: "?q \<noteq> 0" | |
| 735 | using maxge Suc_n_not_le_n le_trans by force | |
| 68076 | 736 | have eq: "(\<Prod>i\<le>n + Max N. f (Suc (i + Max ?Z))) = ?q" for n | 
| 737 | proof - | |
| 738 |       have "(\<Prod>i\<le>n + Max N. f (Suc (i + Max ?Z))) = prod f {Suc (Max ?Z)..n + Max N + Suc (Max ?Z)}" 
 | |
| 739 | proof (rule prod.reindex_cong [where l = "\<lambda>i. i + Suc (Max ?Z)", THEN sym]) | |
| 740 |         show "{Suc (Max ?Z)..n + Max N + Suc (Max ?Z)} = (\<lambda>i. i + Suc (Max ?Z)) ` {..n + Max N}"
 | |
| 741 | using le_Suc_ex by fastforce | |
| 742 | qed (auto simp: inj_on_def) | |
| 68138 | 743 | also have "\<dots> = ?q" | 
| 68076 | 744 | by (rule prod.mono_neutral_right) | 
| 745 | (use Max.coboundedI [OF \<open>finite N\<close>] f in \<open>force+\<close>) | |
| 746 | finally show ?thesis . | |
| 747 | qed | |
| 68361 | 748 | have q: "raw_has_prod f (Suc (Max ?Z)) ?q" | 
| 749 | proof (simp add: raw_has_prod_def) | |
| 68076 | 750 | show "(\<lambda>n. \<Prod>i\<le>n. f (Suc (i + Max ?Z))) \<longlonglongrightarrow> ?q" | 
| 751 | by (rule LIMSEQ_offset[of _ "(Max N)"]) (simp add: eq) | |
| 752 | qed | |
| 68071 | 753 | show ?thesis | 
| 754 | unfolding has_prod_def | |
| 755 | proof (intro disjI2 exI conjI) | |
| 756 | show "prod f N = 0" | |
| 757 | using \<open>f k = 0\<close> \<open>k \<in> N\<close> \<open>finite N\<close> prod_zero by blast | |
| 758 | show "f (Max ?Z) = 0" | |
| 759 | using Max_in [of ?Z] \<open>finite N\<close> \<open>f k = 0\<close> \<open>k \<in> N\<close> by auto | |
| 760 | qed (use q in auto) | |
| 761 | qed | |
| 762 | qed | |
| 763 | ||
| 68651 | 764 | corollary%unimportant has_prod_0: | 
| 68361 | 765 |   fixes f :: "nat \<Rightarrow> 'a::{semidom,t2_space}"
 | 
| 68071 | 766 | assumes "\<And>n. f n = 1" | 
| 767 | shows "f has_prod 1" | |
| 768 | by (simp add: assms has_prod_cong) | |
| 769 | ||
| 68361 | 770 | lemma prodinf_zero[simp]: "prodinf (\<lambda>n. 1::'a::real_normed_field) = 1" | 
| 771 | using has_prod_unique by force | |
| 772 | ||
| 68071 | 773 | lemma convergent_prod_finite: | 
| 774 |   fixes f :: "nat \<Rightarrow> 'a::{idom,t2_space}"
 | |
| 775 | assumes "finite N" "\<And>n. n \<notin> N \<Longrightarrow> f n = 1" | |
| 776 | shows "convergent_prod f" | |
| 777 | proof - | |
| 68361 | 778 | have "\<exists>n p. raw_has_prod f n p" | 
| 68071 | 779 | using assms has_prod_def has_prod_finite by blast | 
| 780 | then show ?thesis | |
| 781 | by (simp add: convergent_prod_def) | |
| 782 | qed | |
| 783 | ||
| 68127 | 784 | lemma has_prod_If_finite_set: | 
| 785 |   fixes f :: "nat \<Rightarrow> 'a::{idom,t2_space}"
 | |
| 786 | shows "finite A \<Longrightarrow> (\<lambda>r. if r \<in> A then f r else 1) has_prod (\<Prod>r\<in>A. f r)" | |
| 787 | using has_prod_finite[of A "(\<lambda>r. if r \<in> A then f r else 1)"] | |
| 788 | by simp | |
| 789 | ||
| 790 | lemma has_prod_If_finite: | |
| 791 |   fixes f :: "nat \<Rightarrow> 'a::{idom,t2_space}"
 | |
| 792 |   shows "finite {r. P r} \<Longrightarrow> (\<lambda>r. if P r then f r else 1) has_prod (\<Prod>r | P r. f r)"
 | |
| 793 |   using has_prod_If_finite_set[of "{r. P r}"] by simp
 | |
| 794 | ||
| 795 | lemma convergent_prod_If_finite_set[simp, intro]: | |
| 796 |   fixes f :: "nat \<Rightarrow> 'a::{idom,t2_space}"
 | |
| 797 | shows "finite A \<Longrightarrow> convergent_prod (\<lambda>r. if r \<in> A then f r else 1)" | |
| 798 | by (simp add: convergent_prod_finite) | |
| 799 | ||
| 800 | lemma convergent_prod_If_finite[simp, intro]: | |
| 801 |   fixes f :: "nat \<Rightarrow> 'a::{idom,t2_space}"
 | |
| 802 |   shows "finite {r. P r} \<Longrightarrow> convergent_prod (\<lambda>r. if P r then f r else 1)"
 | |
| 803 | using convergent_prod_def has_prod_If_finite has_prod_def by fastforce | |
| 804 | ||
| 805 | lemma has_prod_single: | |
| 806 |   fixes f :: "nat \<Rightarrow> 'a::{idom,t2_space}"
 | |
| 807 | shows "(\<lambda>r. if r = i then f r else 1) has_prod f i" | |
| 808 | using has_prod_If_finite[of "\<lambda>r. r = i"] by simp | |
| 809 | ||
| 68136 | 810 | context | 
| 811 | fixes f :: "nat \<Rightarrow> 'a :: real_normed_field" | |
| 812 | begin | |
| 813 | ||
| 814 | lemma convergent_prod_imp_has_prod: | |
| 815 | assumes "convergent_prod f" | |
| 816 | shows "\<exists>p. f has_prod p" | |
| 817 | proof - | |
| 68361 | 818 | obtain M p where p: "raw_has_prod f M p" | 
| 68136 | 819 | using assms convergent_prod_def by blast | 
| 820 | then have "p \<noteq> 0" | |
| 68361 | 821 | using raw_has_prod_nonzero by blast | 
| 68136 | 822 | with p have fnz: "f i \<noteq> 0" if "i \<ge> M" for i | 
| 68361 | 823 | using raw_has_prod_eq_0 that by blast | 
| 68136 | 824 | define C where "C = (\<Prod>n<M. f n)" | 
| 825 | show ?thesis | |
| 826 | proof (cases "\<forall>n\<le>M. f n \<noteq> 0") | |
| 827 | case True | |
| 828 | then have "C \<noteq> 0" | |
| 829 | by (simp add: C_def) | |
| 830 | then show ?thesis | |
| 831 | by (meson True assms convergent_prod_offset_0 fnz has_prod_def nat_le_linear) | |
| 832 | next | |
| 833 | case False | |
| 834 | let ?N = "GREATEST n. f n = 0" | |
| 835 | have 0: "f ?N = 0" | |
| 836 | using fnz False | |
| 837 | by (metis (mono_tags, lifting) GreatestI_ex_nat nat_le_linear) | |
| 838 | have "f i \<noteq> 0" if "i > ?N" for i | |
| 839 | by (metis (mono_tags, lifting) Greatest_le_nat fnz leD linear that) | |
| 68361 | 840 | then have "\<exists>p. raw_has_prod f (Suc ?N) p" | 
| 68136 | 841 | using assms by (auto simp: intro!: convergent_prod_ignore_nonzero_segment) | 
| 842 | then show ?thesis | |
| 843 | unfolding has_prod_def using 0 by blast | |
| 844 | qed | |
| 845 | qed | |
| 846 | ||
| 847 | lemma convergent_prod_has_prod [intro]: | |
| 848 | shows "convergent_prod f \<Longrightarrow> f has_prod (prodinf f)" | |
| 849 | unfolding prodinf_def | |
| 850 | by (metis convergent_prod_imp_has_prod has_prod_unique theI') | |
| 851 | ||
| 852 | lemma convergent_prod_LIMSEQ: | |
| 853 | shows "convergent_prod f \<Longrightarrow> (\<lambda>n. \<Prod>i\<le>n. f i) \<longlonglongrightarrow> prodinf f" | |
| 854 | by (metis convergent_LIMSEQ_iff convergent_prod_has_prod convergent_prod_imp_convergent | |
| 68361 | 855 | convergent_prod_to_zero_iff raw_has_prod_eq_0 has_prod_def prodinf_eq_lim zero_le) | 
| 68136 | 856 | |
| 68651 | 857 | theorem has_prod_iff: "f has_prod x \<longleftrightarrow> convergent_prod f \<and> prodinf f = x" | 
| 68136 | 858 | proof | 
| 859 | assume "f has_prod x" | |
| 860 | then show "convergent_prod f \<and> prodinf f = x" | |
| 861 | apply safe | |
| 862 | using convergent_prod_def has_prod_def apply blast | |
| 863 | using has_prod_unique by blast | |
| 864 | qed auto | |
| 865 | ||
| 866 | lemma convergent_prod_has_prod_iff: "convergent_prod f \<longleftrightarrow> f has_prod prodinf f" | |
| 867 | by (auto simp: has_prod_iff convergent_prod_has_prod) | |
| 868 | ||
| 869 | lemma prodinf_finite: | |
| 870 | assumes N: "finite N" | |
| 871 | and f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 1" | |
| 872 | shows "prodinf f = (\<Prod>n\<in>N. f n)" | |
| 873 | using has_prod_finite[OF assms, THEN has_prod_unique] by simp | |
| 68127 | 874 | |
| 66277 | 875 | end | 
| 68136 | 876 | |
| 68651 | 877 | subsection%unimportant \<open>Infinite products on ordered topological monoids\<close> | 
| 68361 | 878 | |
| 879 | lemma LIMSEQ_prod_0: | |
| 880 |   fixes f :: "nat \<Rightarrow> 'a::{semidom,topological_space}"
 | |
| 881 | assumes "f i = 0" | |
| 882 |   shows "(\<lambda>n. prod f {..n}) \<longlonglongrightarrow> 0"
 | |
| 883 | proof (subst tendsto_cong) | |
| 884 |   show "\<forall>\<^sub>F n in sequentially. prod f {..n} = 0"
 | |
| 885 | proof | |
| 886 |     show "prod f {..n} = 0" if "n \<ge> i" for n
 | |
| 887 | using that assms by auto | |
| 888 | qed | |
| 889 | qed auto | |
| 890 | ||
| 891 | lemma LIMSEQ_prod_nonneg: | |
| 892 |   fixes f :: "nat \<Rightarrow> 'a::{linordered_semidom,linorder_topology}"
 | |
| 893 |   assumes 0: "\<And>n. 0 \<le> f n" and a: "(\<lambda>n. prod f {..n}) \<longlonglongrightarrow> a"
 | |
| 894 | shows "a \<ge> 0" | |
| 895 | by (simp add: "0" prod_nonneg LIMSEQ_le_const [OF a]) | |
| 896 | ||
| 897 | ||
| 898 | context | |
| 899 |   fixes f :: "nat \<Rightarrow> 'a::{linordered_semidom,linorder_topology}"
 | |
| 900 | begin | |
| 901 | ||
| 902 | lemma has_prod_le: | |
| 903 | assumes f: "f has_prod a" and g: "g has_prod b" and le: "\<And>n. 0 \<le> f n \<and> f n \<le> g n" | |
| 904 | shows "a \<le> b" | |
| 905 | proof (cases "a=0 \<or> b=0") | |
| 906 | case True | |
| 907 | then show ?thesis | |
| 908 | proof | |
| 909 | assume [simp]: "a=0" | |
| 910 | have "b \<ge> 0" | |
| 911 | proof (rule LIMSEQ_prod_nonneg) | |
| 912 |       show "(\<lambda>n. prod g {..n}) \<longlonglongrightarrow> b"
 | |
| 913 | using g by (auto simp: has_prod_def raw_has_prod_def LIMSEQ_prod_0) | |
| 914 | qed (use le order_trans in auto) | |
| 915 | then show ?thesis | |
| 916 | by auto | |
| 917 | next | |
| 918 | assume [simp]: "b=0" | |
| 919 | then obtain i where "g i = 0" | |
| 920 | using g by (auto simp: prod_defs) | |
| 921 | then have "f i = 0" | |
| 922 | using antisym le by force | |
| 923 | then have "a=0" | |
| 924 | using f by (auto simp: prod_defs LIMSEQ_prod_0 LIMSEQ_unique) | |
| 925 | then show ?thesis | |
| 926 | by auto | |
| 927 | qed | |
| 928 | next | |
| 929 | case False | |
| 930 | then show ?thesis | |
| 931 | using assms | |
| 932 | unfolding has_prod_def raw_has_prod_def | |
| 933 | by (force simp: LIMSEQ_prod_0 intro!: LIMSEQ_le prod_mono) | |
| 934 | qed | |
| 935 | ||
| 936 | lemma prodinf_le: | |
| 937 | assumes f: "f has_prod a" and g: "g has_prod b" and le: "\<And>n. 0 \<le> f n \<and> f n \<le> g n" | |
| 938 | shows "prodinf f \<le> prodinf g" | |
| 939 | using has_prod_le [OF assms] has_prod_unique f g by blast | |
| 940 | ||
| 68136 | 941 | end | 
| 68361 | 942 | |
| 943 | ||
| 944 | lemma prod_le_prodinf: | |
| 945 |   fixes f :: "nat \<Rightarrow> 'a::{linordered_idom,linorder_topology}"
 | |
| 946 | assumes "f has_prod a" "\<And>i. 0 \<le> f i" "\<And>i. i\<ge>n \<Longrightarrow> 1 \<le> f i" | |
| 947 |   shows "prod f {..<n} \<le> prodinf f"
 | |
| 948 | by(rule has_prod_le[OF has_prod_If_finite_set]) (use assms has_prod_unique in auto) | |
| 949 | ||
| 950 | lemma prodinf_nonneg: | |
| 951 |   fixes f :: "nat \<Rightarrow> 'a::{linordered_idom,linorder_topology}"
 | |
| 952 | assumes "f has_prod a" "\<And>i. 1 \<le> f i" | |
| 953 | shows "1 \<le> prodinf f" | |
| 954 | using prod_le_prodinf[of f a 0] assms | |
| 955 | by (metis order_trans prod_ge_1 zero_le_one) | |
| 956 | ||
| 957 | lemma prodinf_le_const: | |
| 958 | fixes f :: "nat \<Rightarrow> real" | |
| 959 |   assumes "convergent_prod f" "\<And>n. prod f {..<n} \<le> x" 
 | |
| 960 | shows "prodinf f \<le> x" | |
| 961 | by (metis lessThan_Suc_atMost assms convergent_prod_LIMSEQ LIMSEQ_le_const2) | |
| 962 | ||
| 963 | lemma prodinf_eq_one_iff: | |
| 964 | fixes f :: "nat \<Rightarrow> real" | |
| 965 | assumes f: "convergent_prod f" and ge1: "\<And>n. 1 \<le> f n" | |
| 966 | shows "prodinf f = 1 \<longleftrightarrow> (\<forall>n. f n = 1)" | |
| 967 | proof | |
| 968 | assume "prodinf f = 1" | |
| 969 | then have "(\<lambda>n. \<Prod>i<n. f i) \<longlonglongrightarrow> 1" | |
| 970 | using convergent_prod_LIMSEQ[of f] assms by (simp add: LIMSEQ_lessThan_iff_atMost) | |
| 971 |   then have "\<And>i. (\<Prod>n\<in>{i}. f n) \<le> 1"
 | |
| 972 | proof (rule LIMSEQ_le_const) | |
| 973 | have "1 \<le> prod f n" for n | |
| 974 | by (simp add: ge1 prod_ge_1) | |
| 975 |     have "prod f {..<n} = 1" for n
 | |
| 976 | by (metis \<open>\<And>n. 1 \<le> prod f n\<close> \<open>prodinf f = 1\<close> antisym f convergent_prod_has_prod ge1 order_trans prod_le_prodinf zero_le_one) | |
| 977 |     then have "(\<Prod>n\<in>{i}. f n) \<le> prod f {..<n}" if "n \<ge> Suc i" for i n
 | |
| 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: 
69565diff
changeset | 978 | by (metis mult.left_neutral order_refl prod.cong prod.neutral_const prod.lessThan_Suc) | 
| 68361 | 979 |     then show "\<exists>N. \<forall>n\<ge>N. (\<Prod>n\<in>{i}. f n) \<le> prod f {..<n}" for i
 | 
| 980 | by blast | |
| 981 | qed | |
| 982 | with ge1 show "\<forall>n. f n = 1" | |
| 983 | by (auto intro!: antisym) | |
| 984 | qed (metis prodinf_zero fun_eq_iff) | |
| 985 | ||
| 986 | lemma prodinf_pos_iff: | |
| 987 | fixes f :: "nat \<Rightarrow> real" | |
| 988 | assumes "convergent_prod f" "\<And>n. 1 \<le> f n" | |
| 989 | shows "1 < prodinf f \<longleftrightarrow> (\<exists>i. 1 < f i)" | |
| 990 | using prod_le_prodinf[of f 1] prodinf_eq_one_iff | |
| 991 | by (metis convergent_prod_has_prod assms less_le prodinf_nonneg) | |
| 992 | ||
| 993 | lemma less_1_prodinf2: | |
| 994 | fixes f :: "nat \<Rightarrow> real" | |
| 995 | assumes "convergent_prod f" "\<And>n. 1 \<le> f n" "1 < f i" | |
| 996 | shows "1 < prodinf f" | |
| 997 | proof - | |
| 998 | have "1 < (\<Prod>n<Suc i. f n)" | |
| 999 | using assms by (intro less_1_prod2[where i=i]) auto | |
| 1000 | also have "\<dots> \<le> prodinf f" | |
| 1001 | by (intro prod_le_prodinf) (use assms order_trans zero_le_one in \<open>blast+\<close>) | |
| 1002 | finally show ?thesis . | |
| 1003 | qed | |
| 1004 | ||
| 1005 | lemma less_1_prodinf: | |
| 1006 | fixes f :: "nat \<Rightarrow> real" | |
| 1007 | shows "\<lbrakk>convergent_prod f; \<And>n. 1 < f n\<rbrakk> \<Longrightarrow> 1 < prodinf f" | |
| 1008 | by (intro less_1_prodinf2[where i=1]) (auto intro: less_imp_le) | |
| 1009 | ||
| 1010 | lemma prodinf_nonzero: | |
| 1011 |   fixes f :: "nat \<Rightarrow> 'a :: {idom,topological_semigroup_mult,t2_space}"
 | |
| 1012 | assumes "convergent_prod f" "\<And>i. f i \<noteq> 0" | |
| 1013 | shows "prodinf f \<noteq> 0" | |
| 1014 | by (metis assms convergent_prod_offset_0 has_prod_unique raw_has_prod_def has_prod_def) | |
| 1015 | ||
| 1016 | lemma less_0_prodinf: | |
| 1017 | fixes f :: "nat \<Rightarrow> real" | |
| 1018 | assumes f: "convergent_prod f" and 0: "\<And>i. f i > 0" | |
| 1019 | shows "0 < prodinf f" | |
| 1020 | proof - | |
| 1021 | have "prodinf f \<noteq> 0" | |
| 1022 | by (metis assms less_irrefl prodinf_nonzero) | |
| 1023 | moreover have "0 < (\<Prod>n<i. f n)" for i | |
| 1024 | by (simp add: 0 prod_pos) | |
| 1025 | then have "prodinf f \<ge> 0" | |
| 1026 | using convergent_prod_LIMSEQ [OF f] LIMSEQ_prod_nonneg 0 less_le by blast | |
| 1027 | ultimately show ?thesis | |
| 1028 | by auto | |
| 1029 | qed | |
| 1030 | ||
| 1031 | lemma prod_less_prodinf2: | |
| 1032 | fixes f :: "nat \<Rightarrow> real" | |
| 1033 | assumes f: "convergent_prod f" and 1: "\<And>m. m\<ge>n \<Longrightarrow> 1 \<le> f m" and 0: "\<And>m. 0 < f m" and i: "n \<le> i" "1 < f i" | |
| 1034 |   shows "prod f {..<n} < prodinf f"
 | |
| 1035 | proof - | |
| 1036 |   have "prod f {..<n} \<le> prod f {..<i}"
 | |
| 1037 | by (rule prod_mono2) (use assms less_le in auto) | |
| 1038 |   then have "prod f {..<n} < f i * prod f {..<i}"
 | |
| 1039 |     using mult_less_le_imp_less[of 1 "f i" "prod f {..<n}" "prod f {..<i}"] assms
 | |
| 1040 | by (simp add: prod_pos) | |
| 1041 |   moreover have "prod f {..<Suc i} \<le> prodinf f"
 | |
| 1042 | using prod_le_prodinf[of f _ "Suc i"] | |
| 1043 | by (meson "0" "1" Suc_leD convergent_prod_has_prod f \<open>n \<le> i\<close> le_trans less_eq_real_def) | |
| 1044 | ultimately show ?thesis | |
| 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: 
69565diff
changeset | 1045 | by (metis le_less_trans mult.commute not_le prod.lessThan_Suc) | 
| 68361 | 1046 | qed | 
| 1047 | ||
| 1048 | lemma prod_less_prodinf: | |
| 1049 | fixes f :: "nat \<Rightarrow> real" | |
| 1050 | assumes f: "convergent_prod f" and 1: "\<And>m. m\<ge>n \<Longrightarrow> 1 < f m" and 0: "\<And>m. 0 < f m" | |
| 1051 |   shows "prod f {..<n} < prodinf f"
 | |
| 1052 | by (meson "0" "1" f le_less prod_less_prodinf2) | |
| 1053 | ||
| 1054 | lemma raw_has_prodI_bounded: | |
| 1055 | fixes f :: "nat \<Rightarrow> real" | |
| 1056 | assumes pos: "\<And>n. 1 \<le> f n" | |
| 1057 | and le: "\<And>n. (\<Prod>i<n. f i) \<le> x" | |
| 1058 | shows "\<exists>p. raw_has_prod f 0 p" | |
| 1059 | unfolding raw_has_prod_def add_0_right | |
| 1060 | proof (rule exI LIMSEQ_incseq_SUP conjI)+ | |
| 1061 |   show "bdd_above (range (\<lambda>n. prod f {..n}))"
 | |
| 1062 | by (metis bdd_aboveI2 le lessThan_Suc_atMost) | |
| 1063 |   then have "(SUP i. prod f {..i}) > 0"
 | |
| 1064 | by (metis UNIV_I cSUP_upper less_le_trans pos prod_pos zero_less_one) | |
| 1065 |   then show "(SUP i. prod f {..i}) \<noteq> 0"
 | |
| 1066 | by auto | |
| 1067 |   show "incseq (\<lambda>n. prod f {..n})"
 | |
| 1068 | using pos order_trans [OF zero_le_one] by (auto simp: mono_def intro!: prod_mono2) | |
| 1069 | qed | |
| 1070 | ||
| 1071 | lemma convergent_prodI_nonneg_bounded: | |
| 1072 | fixes f :: "nat \<Rightarrow> real" | |
| 1073 | assumes "\<And>n. 1 \<le> f n" "\<And>n. (\<Prod>i<n. f i) \<le> x" | |
| 1074 | shows "convergent_prod f" | |
| 1075 | using convergent_prod_def raw_has_prodI_bounded [OF assms] by blast | |
| 1076 | ||
| 1077 | ||
| 68651 | 1078 | subsection%unimportant \<open>Infinite products on topological spaces\<close> | 
| 68361 | 1079 | |
| 1080 | context | |
| 1081 |   fixes f g :: "nat \<Rightarrow> 'a::{t2_space,topological_semigroup_mult,idom}"
 | |
| 1082 | begin | |
| 1083 | ||
| 1084 | lemma raw_has_prod_mult: "\<lbrakk>raw_has_prod f M a; raw_has_prod g M b\<rbrakk> \<Longrightarrow> raw_has_prod (\<lambda>n. f n * g n) M (a * b)" | |
| 1085 | by (force simp add: prod.distrib tendsto_mult raw_has_prod_def) | |
| 1086 | ||
| 1087 | lemma has_prod_mult_nz: "\<lbrakk>f has_prod a; g has_prod b; a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. f n * g n) has_prod (a * b)" | |
| 1088 | by (simp add: raw_has_prod_mult has_prod_def) | |
| 1089 | ||
| 1090 | end | |
| 1091 | ||
| 1092 | ||
| 1093 | context | |
| 1094 | fixes f g :: "nat \<Rightarrow> 'a::real_normed_field" | |
| 1095 | begin | |
| 1096 | ||
| 1097 | lemma has_prod_mult: | |
| 1098 | assumes f: "f has_prod a" and g: "g has_prod b" | |
| 1099 | shows "(\<lambda>n. f n * g n) has_prod (a * b)" | |
| 1100 | using f [unfolded has_prod_def] | |
| 1101 | proof (elim disjE exE conjE) | |
| 1102 | assume f0: "raw_has_prod f 0 a" | |
| 1103 | show ?thesis | |
| 1104 | using g [unfolded has_prod_def] | |
| 1105 | proof (elim disjE exE conjE) | |
| 1106 | assume g0: "raw_has_prod g 0 b" | |
| 1107 | with f0 show ?thesis | |
| 1108 | by (force simp add: has_prod_def prod.distrib tendsto_mult raw_has_prod_def) | |
| 1109 | next | |
| 1110 | fix j q | |
| 1111 | assume "b = 0" and "g j = 0" and q: "raw_has_prod g (Suc j) q" | |
| 1112 | obtain p where p: "raw_has_prod f (Suc j) p" | |
| 1113 | using f0 raw_has_prod_ignore_initial_segment by blast | |
| 1114 | then have "Ex (raw_has_prod (\<lambda>n. f n * g n) (Suc j))" | |
| 1115 | using q raw_has_prod_mult by blast | |
| 1116 | then show ?thesis | |
| 1117 | using \<open>b = 0\<close> \<open>g j = 0\<close> has_prod_0_iff by fastforce | |
| 1118 | qed | |
| 1119 | next | |
| 1120 | fix i p | |
| 1121 | assume "a = 0" and "f i = 0" and p: "raw_has_prod f (Suc i) p" | |
| 1122 | show ?thesis | |
| 1123 | using g [unfolded has_prod_def] | |
| 1124 | proof (elim disjE exE conjE) | |
| 1125 | assume g0: "raw_has_prod g 0 b" | |
| 1126 | obtain q where q: "raw_has_prod g (Suc i) q" | |
| 1127 | using g0 raw_has_prod_ignore_initial_segment by blast | |
| 1128 | then have "Ex (raw_has_prod (\<lambda>n. f n * g n) (Suc i))" | |
| 1129 | using raw_has_prod_mult p by blast | |
| 1130 | then show ?thesis | |
| 1131 | using \<open>a = 0\<close> \<open>f i = 0\<close> has_prod_0_iff by fastforce | |
| 1132 | next | |
| 1133 | fix j q | |
| 1134 | assume "b = 0" and "g j = 0" and q: "raw_has_prod g (Suc j) q" | |
| 1135 | obtain p' where p': "raw_has_prod f (Suc (max i j)) p'" | |
| 1136 | by (metis raw_has_prod_ignore_initial_segment max_Suc_Suc max_def p) | |
| 1137 | moreover | |
| 1138 | obtain q' where q': "raw_has_prod g (Suc (max i j)) q'" | |
| 1139 | by (metis raw_has_prod_ignore_initial_segment max.cobounded2 max_Suc_Suc q) | |
| 1140 | ultimately show ?thesis | |
| 1141 | using \<open>b = 0\<close> by (simp add: has_prod_def) (metis \<open>f i = 0\<close> \<open>g j = 0\<close> raw_has_prod_mult max_def) | |
| 1142 | qed | |
| 1143 | qed | |
| 1144 | ||
| 1145 | lemma convergent_prod_mult: | |
| 1146 | assumes f: "convergent_prod f" and g: "convergent_prod g" | |
| 1147 | shows "convergent_prod (\<lambda>n. f n * g n)" | |
| 1148 | unfolding convergent_prod_def | |
| 1149 | proof - | |
| 1150 | obtain M p N q where p: "raw_has_prod f M p" and q: "raw_has_prod g N q" | |
| 1151 | using convergent_prod_def f g by blast+ | |
| 1152 | then obtain p' q' where p': "raw_has_prod f (max M N) p'" and q': "raw_has_prod g (max M N) q'" | |
| 1153 | by (meson raw_has_prod_ignore_initial_segment max.cobounded1 max.cobounded2) | |
| 1154 | then show "\<exists>M p. raw_has_prod (\<lambda>n. f n * g n) M p" | |
| 1155 | using raw_has_prod_mult by blast | |
| 1156 | qed | |
| 1157 | ||
| 1158 | lemma prodinf_mult: "convergent_prod f \<Longrightarrow> convergent_prod g \<Longrightarrow> prodinf f * prodinf g = (\<Prod>n. f n * g n)" | |
| 1159 | by (intro has_prod_unique has_prod_mult convergent_prod_has_prod) | |
| 1160 | ||
| 1161 | end | |
| 1162 | ||
| 1163 | context | |
| 1164 | fixes f :: "'i \<Rightarrow> nat \<Rightarrow> 'a::real_normed_field" | |
| 1165 | and I :: "'i set" | |
| 1166 | begin | |
| 1167 | ||
| 1168 | lemma has_prod_prod: "(\<And>i. i \<in> I \<Longrightarrow> (f i) has_prod (x i)) \<Longrightarrow> (\<lambda>n. \<Prod>i\<in>I. f i n) has_prod (\<Prod>i\<in>I. x i)" | |
| 1169 | by (induct I rule: infinite_finite_induct) (auto intro!: has_prod_mult) | |
| 1170 | ||
| 1171 | lemma prodinf_prod: "(\<And>i. i \<in> I \<Longrightarrow> convergent_prod (f i)) \<Longrightarrow> (\<Prod>n. \<Prod>i\<in>I. f i n) = (\<Prod>i\<in>I. \<Prod>n. f i n)" | |
| 1172 | using has_prod_unique[OF has_prod_prod, OF convergent_prod_has_prod] by simp | |
| 1173 | ||
| 1174 | lemma convergent_prod_prod: "(\<And>i. i \<in> I \<Longrightarrow> convergent_prod (f i)) \<Longrightarrow> convergent_prod (\<lambda>n. \<Prod>i\<in>I. f i n)" | |
| 1175 | using convergent_prod_has_prod_iff has_prod_prod prodinf_prod by force | |
| 1176 | ||
| 1177 | end | |
| 1178 | ||
| 68651 | 1179 | subsection%unimportant \<open>Infinite summability on real normed fields\<close> | 
| 68361 | 1180 | |
| 1181 | context | |
| 1182 | fixes f :: "nat \<Rightarrow> 'a::real_normed_field" | |
| 1183 | begin | |
| 1184 | ||
| 1185 | lemma raw_has_prod_Suc_iff: "raw_has_prod f M (a * f M) \<longleftrightarrow> raw_has_prod (\<lambda>n. f (Suc n)) M a \<and> f M \<noteq> 0" | |
| 1186 | proof - | |
| 1187 | have "raw_has_prod f M (a * f M) \<longleftrightarrow> (\<lambda>i. \<Prod>j\<le>Suc i. f (j+M)) \<longlonglongrightarrow> a * f M \<and> a * f M \<noteq> 0" | |
| 1188 | by (subst LIMSEQ_Suc_iff) (simp add: raw_has_prod_def) | |
| 1189 | also have "\<dots> \<longleftrightarrow> (\<lambda>i. (\<Prod>j\<le>i. f (Suc j + M)) * f M) \<longlonglongrightarrow> a * f M \<and> a * f M \<noteq> 0" | |
| 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: 
69565diff
changeset | 1190 | by (simp add: ac_simps atMost_Suc_eq_insert_0 image_Suc_atMost prod.atLeast1_atMost_eq lessThan_Suc_atMost | 
| 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
69565diff
changeset | 1191 | del: prod.cl_ivl_Suc) | 
| 68361 | 1192 | also have "\<dots> \<longleftrightarrow> raw_has_prod (\<lambda>n. f (Suc n)) M a \<and> f M \<noteq> 0" | 
| 1193 | proof safe | |
| 1194 | assume tends: "(\<lambda>i. (\<Prod>j\<le>i. f (Suc j + M)) * f M) \<longlonglongrightarrow> a * f M" and 0: "a * f M \<noteq> 0" | |
| 1195 | with tendsto_divide[OF tends tendsto_const, of "f M"] | |
| 1196 | show "raw_has_prod (\<lambda>n. f (Suc n)) M a" | |
| 1197 | by (simp add: raw_has_prod_def) | |
| 1198 | qed (auto intro: tendsto_mult_right simp: raw_has_prod_def) | |
| 1199 | finally show ?thesis . | |
| 1200 | qed | |
| 1201 | ||
| 1202 | lemma has_prod_Suc_iff: | |
| 1203 | assumes "f 0 \<noteq> 0" shows "(\<lambda>n. f (Suc n)) has_prod a \<longleftrightarrow> f has_prod (a * f 0)" | |
| 1204 | proof (cases "a = 0") | |
| 1205 | case True | |
| 1206 | then show ?thesis | |
| 1207 | proof (simp add: has_prod_def, safe) | |
| 1208 | fix i x | |
| 1209 | assume "f (Suc i) = 0" and "raw_has_prod (\<lambda>n. f (Suc n)) (Suc i) x" | |
| 1210 | then obtain y where "raw_has_prod f (Suc (Suc i)) y" | |
| 1211 | by (metis (no_types) raw_has_prod_eq_0 Suc_n_not_le_n raw_has_prod_Suc_iff raw_has_prod_ignore_initial_segment raw_has_prod_nonzero linear) | |
| 1212 | then show "\<exists>i. f i = 0 \<and> Ex (raw_has_prod f (Suc i))" | |
| 1213 | using \<open>f (Suc i) = 0\<close> by blast | |
| 1214 | next | |
| 1215 | fix i x | |
| 1216 | assume "f i = 0" and x: "raw_has_prod f (Suc i) x" | |
| 1217 | then obtain j where j: "i = Suc j" | |
| 1218 | by (metis assms not0_implies_Suc) | |
| 1219 | moreover have "\<exists> y. raw_has_prod (\<lambda>n. f (Suc n)) i y" | |
| 1220 | using x by (auto simp: raw_has_prod_def) | |
| 1221 | then show "\<exists>i. f (Suc i) = 0 \<and> Ex (raw_has_prod (\<lambda>n. f (Suc n)) (Suc i))" | |
| 1222 | using \<open>f i = 0\<close> j by blast | |
| 1223 | qed | |
| 1224 | next | |
| 1225 | case False | |
| 1226 | then show ?thesis | |
| 1227 | by (auto simp: has_prod_def raw_has_prod_Suc_iff assms) | |
| 1228 | qed | |
| 1229 | ||
| 1230 | lemma convergent_prod_Suc_iff: | |
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1231 | shows "convergent_prod (\<lambda>n. f (Suc n)) = convergent_prod f" | 
| 68361 | 1232 | proof | 
| 1233 | assume "convergent_prod f" | |
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1234 | then obtain M L where M_nz:"\<forall>n\<ge>M. f n \<noteq> 0" and | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1235 | M_L:"(\<lambda>n. \<Prod>i\<le>n. f (i + M)) \<longlonglongrightarrow> L" and "L \<noteq> 0" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1236 | unfolding convergent_prod_altdef by auto | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1237 | have "(\<lambda>n. \<Prod>i\<le>n. f (Suc (i + M))) \<longlonglongrightarrow> L / f M" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1238 | proof - | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1239 |     have "(\<lambda>n. \<Prod>i\<in>{0..Suc n}. f (i + M)) \<longlonglongrightarrow> L"
 | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1240 | using M_L | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1241 | apply (subst (asm) LIMSEQ_Suc_iff[symmetric]) | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1242 | using atLeast0AtMost by auto | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1243 |     then have "(\<lambda>n. f M * (\<Prod>i\<in>{0..n}. f (Suc (i + M)))) \<longlonglongrightarrow> L"
 | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1244 | apply (subst (asm) prod.atLeast0_atMost_Suc_shift) | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1245 | by simp | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1246 |     then have "(\<lambda>n. (\<Prod>i\<in>{0..n}. f (Suc (i + M)))) \<longlonglongrightarrow> L/f M"
 | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1247 | apply (drule_tac tendsto_divide) | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1248 | using M_nz[rule_format,of M,simplified] by auto | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1249 | then show ?thesis unfolding atLeast0AtMost . | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1250 | qed | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1251 | then show "convergent_prod (\<lambda>n. f (Suc n))" unfolding convergent_prod_altdef | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1252 | apply (rule_tac exI[where x=M]) | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1253 | apply (rule_tac exI[where x="L/f M"]) | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1254 | using M_nz \<open>L\<noteq>0\<close> by auto | 
| 68361 | 1255 | next | 
| 1256 | assume "convergent_prod (\<lambda>n. f (Suc n))" | |
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1257 | then obtain M where "\<exists>L. (\<forall>n\<ge>M. f (Suc n) \<noteq> 0) \<and> (\<lambda>n. \<Prod>i\<le>n. f (Suc (i + M))) \<longlonglongrightarrow> L \<and> L \<noteq> 0" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1258 | unfolding convergent_prod_altdef by auto | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1259 | then show "convergent_prod f" unfolding convergent_prod_altdef | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1260 | apply (rule_tac exI[where x="Suc M"]) | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1261 | using Suc_le_D by auto | 
| 68361 | 1262 | qed | 
| 1263 | ||
| 1264 | lemma raw_has_prod_inverse: | |
| 1265 | assumes "raw_has_prod f M a" shows "raw_has_prod (\<lambda>n. inverse (f n)) M (inverse a)" | |
| 1266 | using assms unfolding raw_has_prod_def by (auto dest: tendsto_inverse simp: prod_inversef [symmetric]) | |
| 1267 | ||
| 1268 | lemma has_prod_inverse: | |
| 1269 | assumes "f has_prod a" shows "(\<lambda>n. inverse (f n)) has_prod (inverse a)" | |
| 1270 | using assms raw_has_prod_inverse unfolding has_prod_def by auto | |
| 1271 | ||
| 1272 | lemma convergent_prod_inverse: | |
| 1273 | assumes "convergent_prod f" | |
| 1274 | shows "convergent_prod (\<lambda>n. inverse (f n))" | |
| 1275 | using assms unfolding convergent_prod_def by (blast intro: raw_has_prod_inverse elim: ) | |
| 1276 | ||
| 1277 | end | |
| 1278 | ||
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1279 | context | 
| 68361 | 1280 | fixes f :: "nat \<Rightarrow> 'a::real_normed_field" | 
| 1281 | begin | |
| 1282 | ||
| 1283 | lemma raw_has_prod_Suc_iff': "raw_has_prod f M a \<longleftrightarrow> raw_has_prod (\<lambda>n. f (Suc n)) M (a / f M) \<and> f M \<noteq> 0" | |
| 1284 | by (metis raw_has_prod_eq_0 add.commute add.left_neutral raw_has_prod_Suc_iff raw_has_prod_nonzero le_add1 nonzero_mult_div_cancel_right times_divide_eq_left) | |
| 1285 | ||
| 1286 | lemma has_prod_divide: "f has_prod a \<Longrightarrow> g has_prod b \<Longrightarrow> (\<lambda>n. f n / g n) has_prod (a / b)" | |
| 1287 | unfolding divide_inverse by (intro has_prod_inverse has_prod_mult) | |
| 1288 | ||
| 1289 | lemma convergent_prod_divide: | |
| 1290 | assumes f: "convergent_prod f" and g: "convergent_prod g" | |
| 1291 | shows "convergent_prod (\<lambda>n. f n / g n)" | |
| 1292 | using f g has_prod_divide has_prod_iff by blast | |
| 1293 | ||
| 1294 | lemma prodinf_divide: "convergent_prod f \<Longrightarrow> convergent_prod g \<Longrightarrow> prodinf f / prodinf g = (\<Prod>n. f n / g n)" | |
| 1295 | by (intro has_prod_unique has_prod_divide convergent_prod_has_prod) | |
| 1296 | ||
| 1297 | lemma prodinf_inverse: "convergent_prod f \<Longrightarrow> (\<Prod>n. inverse (f n)) = inverse (\<Prod>n. f n)" | |
| 1298 | by (intro has_prod_unique [symmetric] has_prod_inverse convergent_prod_has_prod) | |
| 1299 | ||
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1300 | lemma has_prod_Suc_imp: | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1301 | assumes "(\<lambda>n. f (Suc n)) has_prod a" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1302 | shows "f has_prod (a * f 0)" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1303 | proof - | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1304 | have "f has_prod (a * f 0)" when "raw_has_prod (\<lambda>n. f (Suc n)) 0 a" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1305 | apply (cases "f 0=0") | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1306 | using that unfolding has_prod_def raw_has_prod_Suc | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1307 | by (auto simp add: raw_has_prod_Suc_iff) | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1308 | moreover have "f has_prod (a * f 0)" when | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1309 | "(\<exists>i q. a = 0 \<and> f (Suc i) = 0 \<and> raw_has_prod (\<lambda>n. f (Suc n)) (Suc i) q)" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1310 | proof - | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1311 | from that | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1312 | obtain i q where "a = 0" "f (Suc i) = 0" "raw_has_prod (\<lambda>n. f (Suc n)) (Suc i) q" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1313 | by auto | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1314 | then show ?thesis unfolding has_prod_def | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1315 | by (auto intro!:exI[where x="Suc i"] simp:raw_has_prod_Suc) | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1316 | qed | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1317 | ultimately show "f has_prod (a * f 0)" using assms unfolding has_prod_def by auto | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1318 | qed | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1319 | |
| 68361 | 1320 | lemma has_prod_iff_shift: | 
| 1321 | assumes "\<And>i. i < n \<Longrightarrow> f i \<noteq> 0" | |
| 1322 | shows "(\<lambda>i. f (i + n)) has_prod a \<longleftrightarrow> f has_prod (a * (\<Prod>i<n. f i))" | |
| 1323 | using assms | |
| 1324 | proof (induct n arbitrary: a) | |
| 1325 | case 0 | |
| 1326 | then show ?case by simp | |
| 1327 | next | |
| 1328 | case (Suc n) | |
| 1329 | then have "(\<lambda>i. f (Suc i + n)) has_prod a \<longleftrightarrow> (\<lambda>i. f (i + n)) has_prod (a * f n)" | |
| 1330 | by (subst has_prod_Suc_iff) auto | |
| 1331 | with Suc show ?case | |
| 1332 | by (simp add: ac_simps) | |
| 1333 | qed | |
| 1334 | ||
| 68651 | 1335 | corollary%unimportant has_prod_iff_shift': | 
| 68361 | 1336 | assumes "\<And>i. i < n \<Longrightarrow> f i \<noteq> 0" | 
| 1337 | shows "(\<lambda>i. f (i + n)) has_prod (a / (\<Prod>i<n. f i)) \<longleftrightarrow> f has_prod a" | |
| 1338 | by (simp add: assms has_prod_iff_shift) | |
| 1339 | ||
| 1340 | lemma has_prod_one_iff_shift: | |
| 1341 | assumes "\<And>i. i < n \<Longrightarrow> f i = 1" | |
| 1342 | shows "(\<lambda>i. f (i+n)) has_prod a \<longleftrightarrow> (\<lambda>i. f i) has_prod a" | |
| 1343 | by (simp add: assms has_prod_iff_shift) | |
| 1344 | ||
| 1345 | lemma convergent_prod_iff_shift: | |
| 1346 | shows "convergent_prod (\<lambda>i. f (i + n)) \<longleftrightarrow> convergent_prod f" | |
| 1347 | apply safe | |
| 1348 | using convergent_prod_offset apply blast | |
| 1349 | using convergent_prod_ignore_initial_segment convergent_prod_def by blast | |
| 1350 | ||
| 1351 | lemma has_prod_split_initial_segment: | |
| 1352 | assumes "f has_prod a" "\<And>i. i < n \<Longrightarrow> f i \<noteq> 0" | |
| 1353 | shows "(\<lambda>i. f (i + n)) has_prod (a / (\<Prod>i<n. f i))" | |
| 1354 | using assms has_prod_iff_shift' by blast | |
| 1355 | ||
| 1356 | lemma prodinf_divide_initial_segment: | |
| 1357 | assumes "convergent_prod f" "\<And>i. i < n \<Longrightarrow> f i \<noteq> 0" | |
| 1358 | shows "(\<Prod>i. f (i + n)) = (\<Prod>i. f i) / (\<Prod>i<n. f i)" | |
| 1359 | by (rule has_prod_unique[symmetric]) (auto simp: assms has_prod_iff_shift) | |
| 1360 | ||
| 1361 | lemma prodinf_split_initial_segment: | |
| 1362 | assumes "convergent_prod f" "\<And>i. i < n \<Longrightarrow> f i \<noteq> 0" | |
| 1363 | shows "prodinf f = (\<Prod>i. f (i + n)) * (\<Prod>i<n. f i)" | |
| 1364 | by (auto simp add: assms prodinf_divide_initial_segment) | |
| 1365 | ||
| 1366 | lemma prodinf_split_head: | |
| 1367 | assumes "convergent_prod f" "f 0 \<noteq> 0" | |
| 1368 | shows "(\<Prod>n. f (Suc n)) = prodinf f / f 0" | |
| 1369 | using prodinf_split_initial_segment[of 1] assms by simp | |
| 1370 | ||
| 1371 | end | |
| 1372 | ||
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1373 | context | 
| 68361 | 1374 | fixes f :: "nat \<Rightarrow> 'a::real_normed_field" | 
| 1375 | begin | |
| 1376 | ||
| 1377 | lemma convergent_prod_inverse_iff: "convergent_prod (\<lambda>n. inverse (f n)) \<longleftrightarrow> convergent_prod f" | |
| 1378 | by (auto dest: convergent_prod_inverse) | |
| 1379 | ||
| 1380 | lemma convergent_prod_const_iff: | |
| 1381 |   fixes c :: "'a :: {real_normed_field}"
 | |
| 1382 | shows "convergent_prod (\<lambda>_. c) \<longleftrightarrow> c = 1" | |
| 1383 | proof | |
| 1384 | assume "convergent_prod (\<lambda>_. c)" | |
| 1385 | then show "c = 1" | |
| 1386 | using convergent_prod_imp_LIMSEQ LIMSEQ_unique by blast | |
| 1387 | next | |
| 1388 | assume "c = 1" | |
| 1389 | then show "convergent_prod (\<lambda>_. c)" | |
| 1390 | by auto | |
| 1391 | qed | |
| 1392 | ||
| 1393 | lemma has_prod_power: "f has_prod a \<Longrightarrow> (\<lambda>i. f i ^ n) has_prod (a ^ n)" | |
| 1394 | by (induction n) (auto simp: has_prod_mult) | |
| 1395 | ||
| 1396 | lemma convergent_prod_power: "convergent_prod f \<Longrightarrow> convergent_prod (\<lambda>i. f i ^ n)" | |
| 1397 | by (induction n) (auto simp: convergent_prod_mult) | |
| 1398 | ||
| 1399 | lemma prodinf_power: "convergent_prod f \<Longrightarrow> prodinf (\<lambda>i. f i ^ n) = prodinf f ^ n" | |
| 1400 | by (metis has_prod_unique convergent_prod_imp_has_prod has_prod_power) | |
| 1401 | ||
| 1402 | end | |
| 1403 | ||
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1404 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1405 | subsection\<open>Exponentials and logarithms\<close> | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1406 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1407 | context | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1408 |   fixes f :: "nat \<Rightarrow> 'a::{real_normed_field,banach}"
 | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1409 | begin | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1410 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1411 | lemma sums_imp_has_prod_exp: | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1412 | assumes "f sums s" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1413 | shows "raw_has_prod (\<lambda>i. exp (f i)) 0 (exp s)" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1414 | using assms continuous_on_exp [of UNIV "\<lambda>x::'a. x"] | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1415 |   using continuous_on_tendsto_compose [of UNIV exp "(\<lambda>n. sum f {..n})" s]
 | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1416 | by (simp add: prod_defs sums_def_le exp_sum) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1417 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1418 | lemma convergent_prod_exp: | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1419 | assumes "summable f" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1420 | shows "convergent_prod (\<lambda>i. exp (f i))" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1421 | using sums_imp_has_prod_exp assms unfolding summable_def convergent_prod_def by blast | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1422 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1423 | lemma prodinf_exp: | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1424 | assumes "summable f" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1425 | shows "prodinf (\<lambda>i. exp (f i)) = exp (suminf f)" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1426 | proof - | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1427 | have "f sums suminf f" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1428 | using assms by blast | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1429 | then have "(\<lambda>i. exp (f i)) has_prod exp (suminf f)" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1430 | by (simp add: has_prod_def sums_imp_has_prod_exp) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1431 | then show ?thesis | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1432 | by (rule has_prod_unique [symmetric]) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1433 | qed | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1434 | |
| 68361 | 1435 | end | 
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1436 | |
| 68651 | 1437 | theorem convergent_prod_iff_summable_real: | 
| 68585 | 1438 | fixes a :: "nat \<Rightarrow> real" | 
| 1439 | assumes "\<And>n. a n > 0" | |
| 1440 | shows "convergent_prod (\<lambda>k. 1 + a k) \<longleftrightarrow> summable a" (is "?lhs = ?rhs") | |
| 1441 | proof | |
| 1442 | assume ?lhs | |
| 1443 | then obtain p where "raw_has_prod (\<lambda>k. 1 + a k) 0 p" | |
| 1444 | by (metis assms add_less_same_cancel2 convergent_prod_offset_0 not_one_less_zero) | |
| 1445 | then have to_p: "(\<lambda>n. \<Prod>k\<le>n. 1 + a k) \<longlonglongrightarrow> p" | |
| 1446 | by (auto simp: raw_has_prod_def) | |
| 1447 | moreover have le: "(\<Sum>k\<le>n. a k) \<le> (\<Prod>k\<le>n. 1 + a k)" for n | |
| 1448 | by (rule sum_le_prod) (use assms less_le in force) | |
| 1449 | have "(\<Prod>k\<le>n. 1 + a k) \<le> p" for n | |
| 1450 | proof (rule incseq_le [OF _ to_p]) | |
| 1451 | show "incseq (\<lambda>n. \<Prod>k\<le>n. 1 + a k)" | |
| 1452 | using assms by (auto simp: mono_def order.strict_implies_order intro!: prod_mono2) | |
| 1453 | qed | |
| 1454 | with le have "(\<Sum>k\<le>n. a k) \<le> p" for n | |
| 1455 | by (metis order_trans) | |
| 1456 | with assms bounded_imp_summable show ?rhs | |
| 1457 | by (metis not_less order.asym) | |
| 1458 | next | |
| 1459 | assume R: ?rhs | |
| 1460 | have "(\<Prod>k\<le>n. 1 + a k) \<le> exp (suminf a)" for n | |
| 1461 | proof - | |
| 1462 | have "(\<Prod>k\<le>n. 1 + a k) \<le> exp (\<Sum>k\<le>n. a k)" for n | |
| 1463 | by (rule prod_le_exp_sum) (use assms less_le in force) | |
| 1464 | moreover have "exp (\<Sum>k\<le>n. a k) \<le> exp (suminf a)" for n | |
| 1465 | unfolding exp_le_cancel_iff | |
| 1466 | by (meson sum_le_suminf R assms finite_atMost less_eq_real_def) | |
| 1467 | ultimately show ?thesis | |
| 1468 | by (meson order_trans) | |
| 1469 | qed | |
| 1470 | then obtain L where L: "(\<lambda>n. \<Prod>k\<le>n. 1 + a k) \<longlonglongrightarrow> L" | |
| 1471 | by (metis assms bounded_imp_convergent_prod convergent_prod_iff_nz_lim le_add_same_cancel1 le_add_same_cancel2 less_le not_le zero_le_one) | |
| 1472 | moreover have "L \<noteq> 0" | |
| 1473 | proof | |
| 1474 | assume "L = 0" | |
| 1475 | with L have "(\<lambda>n. \<Prod>k\<le>n. 1 + a k) \<longlonglongrightarrow> 0" | |
| 1476 | by simp | |
| 1477 | moreover have "(\<Prod>k\<le>n. 1 + a k) > 1" for n | |
| 1478 | by (simp add: assms less_1_prod) | |
| 1479 | ultimately show False | |
| 1480 | by (meson Lim_bounded2 not_one_le_zero less_imp_le) | |
| 1481 | qed | |
| 1482 | ultimately show ?lhs | |
| 1483 | using assms convergent_prod_iff_nz_lim | |
| 1484 | by (metis add_less_same_cancel1 less_le not_le zero_less_one) | |
| 1485 | qed | |
| 1486 | ||
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1487 | lemma exp_suminf_prodinf_real: | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1488 | fixes f :: "nat \<Rightarrow> real" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1489 | assumes ge0:"\<And>n. f n \<ge> 0" and ac: "abs_convergent_prod (\<lambda>n. exp (f n))" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1490 | shows "prodinf (\<lambda>i. exp (f i)) = exp (suminf f)" | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1491 | proof - | 
| 68517 | 1492 | have "summable f" | 
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1493 | using ac unfolding abs_convergent_prod_conv_summable | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1494 | proof (elim summable_comparison_test') | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1495 | fix n | 
| 68517 | 1496 | have "\<bar>f n\<bar> = f n" | 
| 1497 | by (simp add: ge0) | |
| 1498 | also have "\<dots> \<le> exp (f n) - 1" | |
| 1499 | by (metis diff_diff_add exp_ge_add_one_self ge_iff_diff_ge_0) | |
| 1500 | finally show "norm (f n) \<le> norm (exp (f n) - 1)" | |
| 1501 | by simp | |
| 68452 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1502 | qed | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1503 | then show ?thesis | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1504 | by (simp add: prodinf_exp) | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1505 | qed | 
| 
c027dfbfad30
more on infinite products. Also subgroup_imp_subset -> subgroup.subset
 paulson <lp15@cam.ac.uk> parents: 
68426diff
changeset | 1506 | |
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1507 | lemma has_prod_imp_sums_ln_real: | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1508 | fixes f :: "nat \<Rightarrow> real" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1509 | assumes "raw_has_prod f 0 p" and 0: "\<And>x. f x > 0" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1510 | shows "(\<lambda>i. ln (f i)) sums (ln p)" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1511 | proof - | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1512 | have "p > 0" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1513 | using assms unfolding prod_defs by (metis LIMSEQ_prod_nonneg less_eq_real_def) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1514 | then show ?thesis | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1515 |   using assms continuous_on_ln [of "{0<..}" "\<lambda>x. x"]
 | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1516 |   using continuous_on_tendsto_compose [of "{0<..}" ln "(\<lambda>n. prod f {..n})" p]
 | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1517 | by (auto simp: prod_defs sums_def_le ln_prod order_tendstoD) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1518 | qed | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1519 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1520 | lemma summable_ln_real: | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1521 | fixes f :: "nat \<Rightarrow> real" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1522 | assumes f: "convergent_prod f" and 0: "\<And>x. f x > 0" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1523 | shows "summable (\<lambda>i. ln (f i))" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1524 | proof - | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1525 | obtain M p where "raw_has_prod f M p" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1526 | using f convergent_prod_def by blast | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1527 | then consider i where "i<M" "f i = 0" | p where "raw_has_prod f 0 p" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1528 | using raw_has_prod_cases by blast | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1529 | then show ?thesis | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1530 | proof cases | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1531 | case 1 | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1532 | with 0 show ?thesis | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1533 | by (metis less_irrefl) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1534 | next | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1535 | case 2 | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1536 | then show ?thesis | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1537 | using "0" has_prod_imp_sums_ln_real summable_def by blast | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1538 | qed | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1539 | qed | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1540 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1541 | lemma suminf_ln_real: | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1542 | fixes f :: "nat \<Rightarrow> real" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1543 | assumes f: "convergent_prod f" and 0: "\<And>x. f x > 0" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1544 | shows "suminf (\<lambda>i. ln (f i)) = ln (prodinf f)" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1545 | proof - | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1546 | have "f has_prod prodinf f" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1547 | by (simp add: f has_prod_iff) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1548 | then have "raw_has_prod f 0 (prodinf f)" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1549 | by (metis "0" has_prod_def less_irrefl) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1550 | then have "(\<lambda>i. ln (f i)) sums ln (prodinf f)" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1551 | using "0" has_prod_imp_sums_ln_real by blast | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1552 | then show ?thesis | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1553 | by (rule sums_unique [symmetric]) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1554 | qed | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1555 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1556 | lemma prodinf_exp_real: | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1557 | fixes f :: "nat \<Rightarrow> real" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1558 | assumes f: "convergent_prod f" and 0: "\<And>x. f x > 0" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1559 | shows "prodinf f = exp (suminf (\<lambda>i. ln (f i)))" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1560 | by (simp add: "0" f less_0_prodinf suminf_ln_real) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1561 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1562 | |
| 68651 | 1563 | theorem Ln_prodinf_complex: | 
| 68585 | 1564 | fixes z :: "nat \<Rightarrow> complex" | 
| 1565 | assumes z: "\<And>j. z j \<noteq> 0" and \<xi>: "\<xi> \<noteq> 0" | |
| 1566 | shows "((\<lambda>n. \<Prod>j\<le>n. z j) \<longlonglongrightarrow> \<xi>) \<longleftrightarrow> (\<exists>k. (\<lambda>n. (\<Sum>j\<le>n. Ln (z j))) \<longlonglongrightarrow> Ln \<xi> + of_int k * (of_real(2*pi) * \<i>))" (is "?lhs = ?rhs") | |
| 1567 | proof | |
| 1568 | assume L: ?lhs | |
| 1569 | have pnz: "(\<Prod>j\<le>n. z j) \<noteq> 0" for n | |
| 1570 | using z by auto | |
| 1571 | define \<Theta> where "\<Theta> \<equiv> Arg \<xi> + 2*pi" | |
| 1572 | then have "\<Theta> > pi" | |
| 1573 | using Arg_def mpi_less_Im_Ln by fastforce | |
| 1574 | have \<xi>_eq: "\<xi> = cmod \<xi> * exp (\<i> * \<Theta>)" | |
| 1575 | using Arg_def Arg_eq \<xi> unfolding \<Theta>_def by (simp add: algebra_simps exp_add) | |
| 1576 |   define \<theta> where "\<theta> \<equiv> \<lambda>n. THE t. is_Arg (\<Prod>j\<le>n. z j) t \<and> t \<in> {\<Theta>-pi<..\<Theta>+pi}"
 | |
| 1577 |   have uniq: "\<exists>!s. is_Arg (\<Prod>j\<le>n. z j) s \<and> s \<in> {\<Theta>-pi<..\<Theta>+pi}" for n
 | |
| 1578 | using Argument_exists_unique [OF pnz] by metis | |
| 1579 |   have \<theta>: "is_Arg (\<Prod>j\<le>n. z j) (\<theta> n)" and \<theta>_interval: "\<theta> n \<in> {\<Theta>-pi<..\<Theta>+pi}" for n
 | |
| 1580 | unfolding \<theta>_def | |
| 1581 | using theI' [OF uniq] by metis+ | |
| 1582 | have \<theta>_pos: "\<And>j. \<theta> j > 0" | |
| 1583 | using \<theta>_interval \<open>\<Theta> > pi\<close> by simp (meson diff_gt_0_iff_gt less_trans) | |
| 1584 | have "(\<Prod>j\<le>n. z j) = cmod (\<Prod>j\<le>n. z j) * exp (\<i> * \<theta> n)" for n | |
| 1585 | using \<theta> by (auto simp: is_Arg_def) | |
| 1586 | then have eq: "(\<lambda>n. \<Prod>j\<le>n. z j) = (\<lambda>n. cmod (\<Prod>j\<le>n. z j) * exp (\<i> * \<theta> n))" | |
| 1587 | by simp | |
| 1588 | then have "(\<lambda>n. (cmod (\<Prod>j\<le>n. z j)) * exp (\<i> * (\<theta> n))) \<longlonglongrightarrow> \<xi>" | |
| 1589 | using L by force | |
| 1590 | then obtain k where k: "(\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>" | |
| 1591 | using L by (subst (asm) \<xi>_eq) (auto simp add: eq z \<xi> polar_convergence) | |
| 1592 | moreover have "\<forall>\<^sub>F n in sequentially. k n = 0" | |
| 1593 | proof - | |
| 1594 |     have *: "kj = 0" if "dist (vj - real_of_int kj * 2) V < 1" "vj \<in> {V - 1<..V + 1}" for kj vj V
 | |
| 1595 | using that by (auto simp: dist_norm) | |
| 1596 | have "\<forall>\<^sub>F j in sequentially. dist (\<theta> j - of_int (k j) * (2 * pi)) \<Theta> < pi" | |
| 1597 | using tendstoD [OF k] pi_gt_zero by blast | |
| 1598 | then show ?thesis | |
| 1599 | proof (rule eventually_mono) | |
| 1600 | fix j | |
| 1601 | assume d: "dist (\<theta> j - real_of_int (k j) * (2 * pi)) \<Theta> < pi" | |
| 1602 | show "k j = 0" | |
| 1603 | by (rule * [of "\<theta> j/pi" _ "\<Theta>/pi"]) | |
| 1604 | (use \<theta>_interval [of j] d in \<open>simp_all add: divide_simps dist_norm\<close>) | |
| 1605 | qed | |
| 1606 | qed | |
| 1607 | ultimately have \<theta>to\<Theta>: "\<theta> \<longlonglongrightarrow> \<Theta>" | |
| 1608 | apply (simp only: tendsto_def) | |
| 1609 | apply (erule all_forward imp_forward asm_rl)+ | |
| 1610 | apply (drule (1) eventually_conj) | |
| 1611 | apply (auto elim: eventually_mono) | |
| 1612 | done | |
| 1613 | then have to0: "(\<lambda>n. \<bar>\<theta> (Suc n) - \<theta> n\<bar>) \<longlonglongrightarrow> 0" | |
| 1614 | by (metis (full_types) diff_self filterlim_sequentially_Suc tendsto_diff tendsto_rabs_zero) | |
| 1615 | have "\<exists>k. Im (\<Sum>j\<le>n. Ln (z j)) - of_int k * (2*pi) = \<theta> n" for n | |
| 1616 | proof (rule is_Arg_exp_diff_2pi) | |
| 1617 | show "is_Arg (exp (\<Sum>j\<le>n. Ln (z j))) (\<theta> n)" | |
| 1618 | using pnz \<theta> by (simp add: is_Arg_def exp_sum prod_norm) | |
| 1619 | qed | |
| 1620 | then have "\<exists>k. (\<Sum>j\<le>n. Im (Ln (z j))) = \<theta> n + of_int k * (2*pi)" for n | |
| 1621 | by (simp add: algebra_simps) | |
| 1622 | then obtain k where k: "\<And>n. (\<Sum>j\<le>n. Im (Ln (z j))) = \<theta> n + of_int (k n) * (2*pi)" | |
| 1623 | by metis | |
| 1624 | obtain K where "\<forall>\<^sub>F n in sequentially. k n = K" | |
| 1625 | proof - | |
| 1626 | have k_le: "(2*pi) * \<bar>k (Suc n) - k n\<bar> \<le> \<bar>\<theta> (Suc n) - \<theta> n\<bar> + \<bar>Im (Ln (z (Suc n)))\<bar>" for n | |
| 1627 | proof - | |
| 1628 | have "(\<Sum>j\<le>Suc n. Im (Ln (z j))) - (\<Sum>j\<le>n. Im (Ln (z j))) = Im (Ln (z (Suc n)))" | |
| 1629 | by simp | |
| 1630 | then show ?thesis | |
| 1631 | using k [of "Suc n"] k [of n] by (auto simp: abs_if algebra_simps) | |
| 1632 | qed | |
| 1633 | have "z \<longlonglongrightarrow> 1" | |
| 1634 | using L \<xi> convergent_prod_iff_nz_lim z by (blast intro: convergent_prod_imp_LIMSEQ) | |
| 1635 | with z have "(\<lambda>n. Ln (z n)) \<longlonglongrightarrow> Ln 1" | |
| 1636 | using isCont_tendsto_compose [OF continuous_at_Ln] nonpos_Reals_one_I by blast | |
| 1637 | then have "(\<lambda>n. Ln (z n)) \<longlonglongrightarrow> 0" | |
| 1638 | by simp | |
| 1639 | then have "(\<lambda>n. \<bar>Im (Ln (z (Suc n)))\<bar>) \<longlonglongrightarrow> 0" | |
| 1640 | by (metis LIMSEQ_unique \<open>z \<longlonglongrightarrow> 1\<close> continuous_at_Ln filterlim_sequentially_Suc isCont_tendsto_compose nonpos_Reals_one_I tendsto_Im tendsto_rabs_zero_iff zero_complex.simps(2)) | |
| 1641 | then have "\<forall>\<^sub>F n in sequentially. \<bar>Im (Ln (z (Suc n)))\<bar> < 1" | |
| 1642 | by (simp add: order_tendsto_iff) | |
| 1643 | moreover have "\<forall>\<^sub>F n in sequentially. \<bar>\<theta> (Suc n) - \<theta> n\<bar> < 1" | |
| 1644 | using to0 by (simp add: order_tendsto_iff) | |
| 1645 | ultimately have "\<forall>\<^sub>F n in sequentially. (2*pi) * \<bar>k (Suc n) - k n\<bar> < 1 + 1" | |
| 1646 | proof (rule eventually_elim2) | |
| 1647 | fix n | |
| 1648 | assume "\<bar>Im (Ln (z (Suc n)))\<bar> < 1" and "\<bar>\<theta> (Suc n) - \<theta> n\<bar> < 1" | |
| 1649 | with k_le [of n] show "2 * pi * real_of_int \<bar>k (Suc n) - k n\<bar> < 1 + 1" | |
| 1650 | by linarith | |
| 1651 | qed | |
| 1652 | then have "\<forall>\<^sub>F n in sequentially. real_of_int\<bar>k (Suc n) - k n\<bar> < 1" | |
| 1653 | proof (rule eventually_mono) | |
| 1654 | fix n :: "nat" | |
| 1655 | assume "2 * pi * \<bar>k (Suc n) - k n\<bar> < 1 + 1" | |
| 1656 | then have "\<bar>k (Suc n) - k n\<bar> < 2 / (2*pi)" | |
| 1657 | by (simp add: field_simps) | |
| 1658 | also have "... < 1" | |
| 1659 | using pi_ge_two by auto | |
| 1660 | finally show "real_of_int \<bar>k (Suc n) - k n\<bar> < 1" . | |
| 1661 | qed | |
| 1662 | then obtain N where N: "\<And>n. n\<ge>N \<Longrightarrow> \<bar>k (Suc n) - k n\<bar> = 0" | |
| 1663 | using eventually_sequentially less_irrefl of_int_abs by fastforce | |
| 1664 | have "k (N+i) = k N" for i | |
| 1665 | proof (induction i) | |
| 1666 | case (Suc i) | |
| 1667 | with N [of "N+i"] show ?case | |
| 1668 | by auto | |
| 1669 | qed simp | |
| 1670 | then have "\<And>n. n\<ge>N \<Longrightarrow> k n = k N" | |
| 1671 | using le_Suc_ex by auto | |
| 1672 | then show ?thesis | |
| 1673 | by (force simp add: eventually_sequentially intro: that) | |
| 1674 | qed | |
| 1675 | with \<theta>to\<Theta> have "(\<lambda>n. (\<Sum>j\<le>n. Im (Ln (z j)))) \<longlonglongrightarrow> \<Theta> + of_int K * (2*pi)" | |
| 1676 | by (simp add: k tendsto_add tendsto_mult Lim_eventually) | |
| 1677 | moreover have "(\<lambda>n. (\<Sum>k\<le>n. Re (Ln (z k)))) \<longlonglongrightarrow> Re (Ln \<xi>)" | |
| 1678 | using assms continuous_imp_tendsto [OF isCont_ln tendsto_norm [OF L]] | |
| 1679 | by (simp add: o_def flip: prod_norm ln_prod) | |
| 1680 | ultimately show ?rhs | |
| 1681 | by (rule_tac x="K+1" in exI) (auto simp: tendsto_complex_iff \<Theta>_def Arg_def assms algebra_simps) | |
| 1682 | next | |
| 1683 | assume ?rhs | |
| 1684 | then obtain r where r: "(\<lambda>n. (\<Sum>k\<le>n. Ln (z k))) \<longlonglongrightarrow> Ln \<xi> + of_int r * (of_real(2*pi) * \<i>)" .. | |
| 1685 | have "(\<lambda>n. exp (\<Sum>k\<le>n. Ln (z k))) \<longlonglongrightarrow> \<xi>" | |
| 1686 | using assms continuous_imp_tendsto [OF isCont_exp r] exp_integer_2pi [of r] | |
| 1687 | by (simp add: o_def exp_add algebra_simps) | |
| 1688 | moreover have "exp (\<Sum>k\<le>n. Ln (z k)) = (\<Prod>k\<le>n. z k)" for n | |
| 1689 | by (simp add: exp_sum add_eq_0_iff assms) | |
| 1690 | ultimately show ?lhs | |
| 1691 | by auto | |
| 1692 | qed | |
| 1693 | ||
| 1694 | text\<open>Prop 17.2 of Bak and Newman, Complex Analysis, p.242\<close> | |
| 1695 | proposition convergent_prod_iff_summable_complex: | |
| 1696 | fixes z :: "nat \<Rightarrow> complex" | |
| 1697 | assumes "\<And>k. z k \<noteq> 0" | |
| 1698 | shows "convergent_prod (\<lambda>k. z k) \<longleftrightarrow> summable (\<lambda>k. Ln (z k))" (is "?lhs = ?rhs") | |
| 1699 | proof | |
| 1700 | assume ?lhs | |
| 1701 | then obtain p where p: "(\<lambda>n. \<Prod>k\<le>n. z k) \<longlonglongrightarrow> p" and "p \<noteq> 0" | |
| 1702 | using convergent_prod_LIMSEQ prodinf_nonzero add_eq_0_iff assms by fastforce | |
| 1703 | then show ?rhs | |
| 1704 | using Ln_prodinf_complex assms | |
| 1705 | by (auto simp: prodinf_nonzero summable_def sums_def_le) | |
| 1706 | next | |
| 1707 | assume R: ?rhs | |
| 1708 | have "(\<Prod>k\<le>n. z k) = exp (\<Sum>k\<le>n. Ln (z k))" for n | |
| 1709 | by (simp add: exp_sum add_eq_0_iff assms) | |
| 1710 | then have "(\<lambda>n. \<Prod>k\<le>n. z k) \<longlonglongrightarrow> exp (suminf (\<lambda>k. Ln (z k)))" | |
| 1711 | using continuous_imp_tendsto [OF isCont_exp summable_LIMSEQ' [OF R]] by (simp add: o_def) | |
| 1712 | then show ?lhs | |
| 1713 | by (subst convergent_prod_iff_convergent) (auto simp: convergent_def tendsto_Lim assms add_eq_0_iff) | |
| 1714 | qed | |
| 1715 | ||
| 68586 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1716 | text\<open>Prop 17.3 of Bak and Newman, Complex Analysis\<close> | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1717 | proposition summable_imp_convergent_prod_complex: | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1718 | fixes z :: "nat \<Rightarrow> complex" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1719 | assumes z: "summable (\<lambda>k. norm (z k))" and non0: "\<And>k. z k \<noteq> -1" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1720 | shows "convergent_prod (\<lambda>k. 1 + z k)" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1721 | proof - | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1722 | note if_cong [cong] power_Suc [simp del] | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1723 | obtain N where N: "\<And>k. k\<ge>N \<Longrightarrow> norm (z k) < 1/2" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1724 | using summable_LIMSEQ_zero [OF z] | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1725 | by (metis diff_zero dist_norm half_gt_zero_iff less_numeral_extra(1) lim_sequentially tendsto_norm_zero_iff) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1726 | have "norm (Ln (1 + z k)) \<le> 2 * norm (z k)" if "k \<ge> N" for k | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1727 | proof (cases "z k = 0") | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1728 | case False | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1729 | let ?f = "\<lambda>i. cmod ((- 1) ^ i * z k ^ i / of_nat (Suc i))" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1730 | have normf: "norm (?f n) \<le> (1 / 2) ^ n" for n | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1731 | proof - | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1732 | have "norm (?f n) = cmod (z k) ^ n / cmod (1 + of_nat n)" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1733 | by (auto simp: norm_divide norm_mult norm_power) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1734 | also have "\<dots> \<le> cmod (z k) ^ n" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1735 | by (auto simp: divide_simps mult_le_cancel_left1 in_Reals_norm) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1736 | also have "\<dots> \<le> (1 / 2) ^ n" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1737 | using N [OF that] by (simp add: power_mono) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1738 | finally show "norm (?f n) \<le> (1 / 2) ^ n" . | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1739 | qed | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1740 | have summablef: "summable ?f" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1741 | by (intro normf summable_comparison_test' [OF summable_geometric [of "1/2"]]) auto | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1742 | have "(\<lambda>n. (- 1) ^ Suc n / of_nat n * z k ^ n) sums Ln (1 + z k)" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1743 | using Ln_series [of "z k"] N that by fastforce | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1744 | then have *: "(\<lambda>i. z k * (((- 1) ^ i * z k ^ i) / (Suc i))) sums Ln (1 + z k)" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1745 | using sums_split_initial_segment [where n= 1] by (force simp: power_Suc mult_ac) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1746 | then have "norm (Ln (1 + z k)) = norm (suminf (\<lambda>i. z k * (((- 1) ^ i * z k ^ i) / (Suc i))))" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1747 | using sums_unique by force | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1748 | also have "\<dots> = norm (z k * suminf (\<lambda>i. ((- 1) ^ i * z k ^ i) / (Suc i)))" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1749 | apply (subst suminf_mult) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1750 | using * False | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1751 | by (auto simp: sums_summable intro: summable_mult_D [of "z k"]) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1752 | also have "\<dots> = norm (z k) * norm (suminf (\<lambda>i. ((- 1) ^ i * z k ^ i) / (Suc i)))" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1753 | by (simp add: norm_mult) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1754 | also have "\<dots> \<le> norm (z k) * suminf (\<lambda>i. norm (((- 1) ^ i * z k ^ i) / (Suc i)))" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1755 | by (intro mult_left_mono summable_norm summablef) auto | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1756 | also have "\<dots> \<le> norm (z k) * suminf (\<lambda>i. (1/2) ^ i)" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1757 | by (intro mult_left_mono suminf_le) (use summable_geometric [of "1/2"] summablef normf in auto) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1758 | also have "\<dots> \<le> norm (z k) * 2" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1759 | using suminf_geometric [of "1/2::real"] by simp | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1760 | finally show ?thesis | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1761 | by (simp add: mult_ac) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1762 | qed simp | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1763 | then have "summable (\<lambda>k. Ln (1 + z k))" | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1764 | by (metis summable_comparison_test summable_mult z) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1765 | with non0 show ?thesis | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1766 | by (simp add: add_eq_0_iff convergent_prod_iff_summable_complex) | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1767 | qed | 
| 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1768 | |
| 68616 | 1769 | lemma summable_Ln_complex: | 
| 1770 | fixes z :: "nat \<Rightarrow> complex" | |
| 1771 | assumes "convergent_prod z" "\<And>k. z k \<noteq> 0" | |
| 1772 | shows "summable (\<lambda>k. Ln (z k))" | |
| 1773 | using convergent_prod_def assms convergent_prod_iff_summable_complex by blast | |
| 1774 | ||
| 68586 
006da53a8ac1
infinite products: the final piece
 paulson <lp15@cam.ac.uk> parents: 
68585diff
changeset | 1775 | |
| 68651 | 1776 | subsection%unimportant \<open>Embeddings from the reals into some complete real normed field\<close> | 
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1777 | |
| 68426 | 1778 | lemma tendsto_eq_of_real_lim: | 
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1779 |   assumes "(\<lambda>n. of_real (f n) :: 'a::{complete_space,real_normed_field}) \<longlonglongrightarrow> q"
 | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1780 | shows "q = of_real (lim f)" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1781 | proof - | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1782 | have "convergent (\<lambda>n. of_real (f n) :: 'a)" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1783 | using assms convergent_def by blast | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1784 | then have "convergent f" | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1785 | unfolding convergent_def | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1786 | by (simp add: convergent_eq_Cauchy Cauchy_def) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1787 | then show ?thesis | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1788 | by (metis LIMSEQ_unique assms convergentD sequentially_bot tendsto_Lim tendsto_of_real) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1789 | qed | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1790 | |
| 68426 | 1791 | lemma tendsto_eq_of_real: | 
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1792 |   assumes "(\<lambda>n. of_real (f n) :: 'a::{complete_space,real_normed_field}) \<longlonglongrightarrow> q"
 | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1793 | obtains r where "q = of_real r" | 
| 68426 | 1794 | using tendsto_eq_of_real_lim assms by blast | 
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1795 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1796 | lemma has_prod_of_real_iff: | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1797 |   "(\<lambda>n. of_real (f n) :: 'a::{complete_space,real_normed_field}) has_prod of_real c \<longleftrightarrow> f has_prod c"
 | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1798 | (is "?lhs = ?rhs") | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1799 | proof | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1800 | assume ?lhs | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1801 | then show ?rhs | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1802 | apply (auto simp: prod_defs LIMSEQ_prod_0 tendsto_of_real_iff simp flip: of_real_prod) | 
| 68426 | 1803 | using tendsto_eq_of_real | 
| 68424 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1804 | by (metis of_real_0 tendsto_of_real_iff) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1805 | next | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1806 | assume ?rhs | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1807 | with tendsto_of_real_iff show ?lhs | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1808 | by (fastforce simp: prod_defs simp flip: of_real_prod) | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1809 | qed | 
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1810 | |
| 
02e5a44ffe7d
the last of the infinite product proofs
 paulson <lp15@cam.ac.uk> parents: 
68361diff
changeset | 1811 | end |