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