| author | Fabian Huch <huch@in.tum.de> | 
| Tue, 11 Jun 2024 14:27:04 +0200 | |
| changeset 80347 | 613ac8c77a84 | 
| parent 79757 | f20ac6788faa | 
| child 80768 | c7723cc15de8 | 
| permissions | -rw-r--r-- | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1 | (* | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2 | Title: HOL/Analysis/Infinite_Sum.thy | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 3 | Author: Dominique Unruh, University of Tartu | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 4 | Manuel Eberl, University of Innsbruck | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 5 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 6 | A theory of sums over possibly infinite sets. | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 7 | *) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 8 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 9 | section \<open>Infinite sums\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 10 | \<^latex>\<open>\label{section:Infinite_Sum}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 11 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 12 | text \<open>In this theory, we introduce the definition of infinite sums, i.e., sums ranging over an | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 13 | infinite, potentially uncountable index set with no particular ordering. | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 14 | (This is different from series. Those are sums indexed by natural numbers, | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 15 | and the order of the index set matters.) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 16 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 17 | Our definition is quite standard: $s:=\sum_{x\in A} f(x)$ is the limit of finite sums $s_F:=\sum_{x\in F} f(x)$ for increasing $F$.
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 18 | That is, $s$ is the limit of the net $s_F$ where $F$ are finite subsets of $A$ ordered by inclusion. | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 19 | We believe that this is the standard definition for such sums. | 
| 76987 | 20 | See, e.g., Definition 4.11 in \<^cite>\<open>"conway2013course"\<close>. | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 21 | This definition is quite general: it is well-defined whenever $f$ takes values in some | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 22 | commutative monoid endowed with a Hausdorff topology. | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 23 | (Examples are reals, complex numbers, normed vector spaces, and more.)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 24 | |
| 76999 
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
 wenzelm parents: 
76988diff
changeset | 25 | theory Infinite_Sum | 
| 
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
 wenzelm parents: 
76988diff
changeset | 26 | imports | 
| 
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
 wenzelm parents: 
76988diff
changeset | 27 | Elementary_Topology | 
| 
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
 wenzelm parents: 
76988diff
changeset | 28 | "HOL-Library.Extended_Nonnegative_Real" | 
| 
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
 wenzelm parents: 
76988diff
changeset | 29 | "HOL-Library.Complex_Order" | 
| 
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
 wenzelm parents: 
76988diff
changeset | 30 | begin | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 31 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 32 | subsection \<open>Definition and syntax\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 33 | |
| 77359 | 34 | definition HAS_SUM :: \<open>('a \<Rightarrow> 'b :: {comm_monoid_add, topological_space}) \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool\<close> 
 | 
| 35 | where has_sum_def: \<open>HAS_SUM f A x \<equiv> (sum f \<longlongrightarrow> x) (finite_subsets_at_top A)\<close> | |
| 36 | ||
| 37 | abbreviation has_sum (infixr "has'_sum" 46) where | |
| 38 | "(f has_sum S) A \<equiv> HAS_SUM f A S" | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 39 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 40 | definition summable_on :: "('a \<Rightarrow> 'b::{comm_monoid_add, topological_space}) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "summable'_on" 46) where
 | 
| 77359 | 41 | "f summable_on A \<equiv> (\<exists>x. (f has_sum x) A)" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 42 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 43 | definition infsum :: "('a \<Rightarrow> 'b::{comm_monoid_add,t2_space}) \<Rightarrow> 'a set \<Rightarrow> 'b" where
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 44 | "infsum f A = (if f summable_on A then Lim (finite_subsets_at_top A) (sum f) else 0)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 45 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 46 | abbreviation abs_summable_on :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "abs'_summable'_on" 46) where
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 47 | "f abs_summable_on A \<equiv> (\<lambda>x. norm (f x)) summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 48 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 49 | syntax (ASCII) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 50 |   "_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add"  ("(3INFSUM (_/:_)./ _)" [0, 51, 10] 10)
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 51 | syntax | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 52 |   "_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add"  ("(2\<Sum>\<^sub>\<infinity>(_/\<in>_)./ _)" [0, 51, 10] 10)
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 53 | translations \<comment> \<open>Beware of argument permutation!\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 54 | "\<Sum>\<^sub>\<infinity>i\<in>A. b" \<rightleftharpoons> "CONST infsum (\<lambda>i. b) A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 55 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 56 | syntax (ASCII) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 57 |   "_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a"  ("(3INFSUM _./ _)" [0, 10] 10)
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 58 | syntax | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 59 |   "_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a"  ("(2\<Sum>\<^sub>\<infinity>_./ _)" [0, 10] 10)
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 60 | translations | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 61 | "\<Sum>\<^sub>\<infinity>x. t" \<rightleftharpoons> "CONST infsum (\<lambda>x. t) (CONST UNIV)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 62 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 63 | syntax (ASCII) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 64 |   "_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(3INFSUM _ |/ _./ _)" [0, 0, 10] 10)
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 65 | syntax | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 66 |   "_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(2\<Sum>\<^sub>\<infinity>_ | (_)./ _)" [0, 0, 10] 10)
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 67 | translations | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 68 |   "\<Sum>\<^sub>\<infinity>x|P. t" => "CONST infsum (\<lambda>x. t) {x. P}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 69 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 70 | print_translation \<open> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 71 | let | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 72 |   fun sum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 73 | if x <> y then raise Match | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 74 | else | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 75 | let | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 76 | val x' = Syntax_Trans.mark_bound_body (x, Tx); | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 77 | val t' = subst_bound (x', t); | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 78 | val P' = subst_bound (x', P); | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 79 | in | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 80 |             Syntax.const @{syntax_const "_qinfsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 81 | end | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 82 | | sum_tr' _ = raise Match; | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 83 | in [(@{const_syntax infsum}, K sum_tr')] end
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 84 | \<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 85 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 86 | subsection \<open>General properties\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 87 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 88 | lemma infsumI: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 89 |   fixes f g :: \<open>'a \<Rightarrow> 'b::{comm_monoid_add, t2_space}\<close>
 | 
| 77359 | 90 | assumes \<open>(f has_sum x) A\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 91 | shows \<open>infsum f A = x\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 92 | by (metis assms finite_subsets_at_top_neq_bot infsum_def summable_on_def has_sum_def tendsto_Lim) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 93 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 94 | lemma infsum_eqI: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 95 |   fixes f g :: \<open>'a \<Rightarrow> 'b::{comm_monoid_add, t2_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 96 | assumes \<open>x = y\<close> | 
| 77359 | 97 | assumes \<open>(f has_sum x) A\<close> | 
| 98 | assumes \<open>(g has_sum y) B\<close> | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 99 | shows \<open>infsum f A = infsum g B\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 100 | using assms infsumI by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 101 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 102 | lemma infsum_eqI': | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 103 |   fixes f g :: \<open>'a \<Rightarrow> 'b::{comm_monoid_add, t2_space}\<close>
 | 
| 77359 | 104 | assumes \<open>\<And>x. (f has_sum x) A \<longleftrightarrow> (g has_sum x) B\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 105 | shows \<open>infsum f A = infsum g B\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 106 | by (metis assms infsum_def infsum_eqI summable_on_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 107 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 108 | lemma infsum_not_exists: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 109 |   fixes f :: \<open>'a \<Rightarrow> 'b::{comm_monoid_add, t2_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 110 | assumes \<open>\<not> f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 111 | shows \<open>infsum f A = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 112 | by (simp add: assms infsum_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 113 | |
| 77359 | 114 | lemma summable_iff_has_sum_infsum: "f summable_on A \<longleftrightarrow> (f has_sum (infsum f A)) A" | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 115 | using infsumI summable_on_def by blast | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 116 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 117 | lemma has_sum_infsum[simp]: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 118 | assumes \<open>f summable_on S\<close> | 
| 77359 | 119 | shows \<open>(f has_sum (infsum f S)) S\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 120 | using assms summable_iff_has_sum_infsum by blast | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 121 | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 122 | lemma has_sum_cong_neutral: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 123 |   fixes f g :: \<open>'a \<Rightarrow> 'b::{comm_monoid_add, topological_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 124 | assumes \<open>\<And>x. x\<in>T-S \<Longrightarrow> g x = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 125 | assumes \<open>\<And>x. x\<in>S-T \<Longrightarrow> f x = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 126 | assumes \<open>\<And>x. x\<in>S\<inter>T \<Longrightarrow> f x = g x\<close> | 
| 77359 | 127 | shows "(f has_sum x) S \<longleftrightarrow> (g has_sum x) T" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 128 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 129 | have \<open>eventually P (filtermap (sum f) (finite_subsets_at_top S)) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 130 | = eventually P (filtermap (sum g) (finite_subsets_at_top T))\<close> for P | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 131 | proof | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 132 | assume \<open>eventually P (filtermap (sum f) (finite_subsets_at_top S))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 133 | then obtain F0 where \<open>finite F0\<close> and \<open>F0 \<subseteq> S\<close> and F0_P: \<open>\<And>F. finite F \<Longrightarrow> F \<subseteq> S \<Longrightarrow> F \<supseteq> F0 \<Longrightarrow> P (sum f F)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 134 | by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 135 | define F0' where \<open>F0' = F0 \<inter> T\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 136 | have [simp]: \<open>finite F0'\<close> \<open>F0' \<subseteq> T\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 137 | by (simp_all add: F0'_def \<open>finite F0\<close>) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 138 | have \<open>P (sum g F)\<close> if \<open>finite F\<close> \<open>F \<subseteq> T\<close> \<open>F \<supseteq> F0'\<close> for F | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 139 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 140 | have \<open>P (sum f ((F\<inter>S) \<union> (F0\<inter>S)))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 141 | by (intro F0_P) (use \<open>F0 \<subseteq> S\<close> \<open>finite F0\<close> that in auto) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 142 | also have \<open>sum f ((F\<inter>S) \<union> (F0\<inter>S)) = sum g F\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 143 | by (intro sum.mono_neutral_cong) (use that \<open>finite F0\<close> F0'_def assms in auto) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 144 | finally show ?thesis . | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 145 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 146 | with \<open>F0' \<subseteq> T\<close> \<open>finite F0'\<close> show \<open>eventually P (filtermap (sum g) (finite_subsets_at_top T))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 147 | by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 148 | next | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 149 | assume \<open>eventually P (filtermap (sum g) (finite_subsets_at_top T))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 150 | then obtain F0 where \<open>finite F0\<close> and \<open>F0 \<subseteq> T\<close> and F0_P: \<open>\<And>F. finite F \<Longrightarrow> F \<subseteq> T \<Longrightarrow> F \<supseteq> F0 \<Longrightarrow> P (sum g F)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 151 | by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 152 | define F0' where \<open>F0' = F0 \<inter> S\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 153 | have [simp]: \<open>finite F0'\<close> \<open>F0' \<subseteq> S\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 154 | by (simp_all add: F0'_def \<open>finite F0\<close>) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 155 | have \<open>P (sum f F)\<close> if \<open>finite F\<close> \<open>F \<subseteq> S\<close> \<open>F \<supseteq> F0'\<close> for F | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 156 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 157 | have \<open>P (sum g ((F\<inter>T) \<union> (F0\<inter>T)))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 158 | by (intro F0_P) (use \<open>F0 \<subseteq> T\<close> \<open>finite F0\<close> that in auto) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 159 | also have \<open>sum g ((F\<inter>T) \<union> (F0\<inter>T)) = sum f F\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 160 | by (intro sum.mono_neutral_cong) (use that \<open>finite F0\<close> F0'_def assms in auto) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 161 | finally show ?thesis . | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 162 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 163 | with \<open>F0' \<subseteq> S\<close> \<open>finite F0'\<close> show \<open>eventually P (filtermap (sum f) (finite_subsets_at_top S))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 164 | by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 165 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 166 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 167 | then have tendsto_x: "(sum f \<longlongrightarrow> x) (finite_subsets_at_top S) \<longleftrightarrow> (sum g \<longlongrightarrow> x) (finite_subsets_at_top T)" for x | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 168 | by (simp add: le_filter_def filterlim_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 169 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 170 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 171 | by (simp add: has_sum_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 172 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 173 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 174 | lemma summable_on_cong_neutral: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 175 |   fixes f g :: \<open>'a \<Rightarrow> 'b::{comm_monoid_add, topological_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 176 | assumes \<open>\<And>x. x\<in>T-S \<Longrightarrow> g x = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 177 | assumes \<open>\<And>x. x\<in>S-T \<Longrightarrow> f x = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 178 | assumes \<open>\<And>x. x\<in>S\<inter>T \<Longrightarrow> f x = g x\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 179 | shows "f summable_on S \<longleftrightarrow> g summable_on T" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 180 | using has_sum_cong_neutral[of T S g f, OF assms] | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 181 | by (simp add: summable_on_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 182 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 183 | lemma infsum_cong_neutral: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 184 |   fixes f g :: \<open>'a \<Rightarrow> 'b::{comm_monoid_add, t2_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 185 | assumes \<open>\<And>x. x\<in>T-S \<Longrightarrow> g x = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 186 | assumes \<open>\<And>x. x\<in>S-T \<Longrightarrow> f x = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 187 | assumes \<open>\<And>x. x\<in>S\<inter>T \<Longrightarrow> f x = g x\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 188 | shows \<open>infsum f S = infsum g T\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 189 | by (smt (verit, best) assms has_sum_cong_neutral infsum_eqI') | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 190 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 191 | lemma has_sum_cong: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 192 | assumes "\<And>x. x\<in>A \<Longrightarrow> f x = g x" | 
| 77359 | 193 | shows "(f has_sum x) A \<longleftrightarrow> (g has_sum x) A" | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 194 | using assms by (intro has_sum_cong_neutral) auto | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 195 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 196 | lemma summable_on_cong: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 197 | assumes "\<And>x. x\<in>A \<Longrightarrow> f x = g x" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 198 | shows "f summable_on A \<longleftrightarrow> g summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 199 | by (metis assms summable_on_def has_sum_cong) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 200 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 201 | lemma infsum_cong: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 202 | assumes "\<And>x. x\<in>A \<Longrightarrow> f x = g x" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 203 | shows "infsum f A = infsum g A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 204 | using assms infsum_eqI' has_sum_cong by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 205 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 206 | lemma summable_on_cofin_subset: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 207 | fixes f :: "'a \<Rightarrow> 'b::topological_ab_group_add" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 208 | assumes "f summable_on A" and [simp]: "finite F" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 209 | shows "f summable_on (A - F)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 210 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 211 | from assms(1) obtain x where lim_f: "(sum f \<longlongrightarrow> x) (finite_subsets_at_top A)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 212 | unfolding summable_on_def has_sum_def by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 213 | define F' where "F' = F\<inter>A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 214 | with assms have "finite F'" and "A-F = A-F'" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 215 | by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 216 | have "filtermap ((\<union>)F') (finite_subsets_at_top (A-F)) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 217 | \<le> finite_subsets_at_top A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 218 | proof (rule filter_leI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 219 | fix P assume "eventually P (finite_subsets_at_top A)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 220 | then obtain X where [simp]: "finite X" and XA: "X \<subseteq> A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 221 | and P: "\<forall>Y. finite Y \<and> X \<subseteq> Y \<and> Y \<subseteq> A \<longrightarrow> P Y" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 222 | unfolding eventually_finite_subsets_at_top by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 223 | define X' where "X' = X-F" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 224 | hence [simp]: "finite X'" and [simp]: "X' \<subseteq> A-F" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 225 | using XA by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 226 | hence "finite Y \<and> X' \<subseteq> Y \<and> Y \<subseteq> A - F \<longrightarrow> P (F' \<union> Y)" for Y | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 227 | using P XA unfolding X'_def using F'_def \<open>finite F'\<close> by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 228 | thus "eventually P (filtermap ((\<union>) F') (finite_subsets_at_top (A - F)))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 229 | unfolding eventually_filtermap eventually_finite_subsets_at_top | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 230 | by (rule_tac x=X' in exI, simp) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 231 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 232 | with lim_f have "(sum f \<longlongrightarrow> x) (filtermap ((\<union>)F') (finite_subsets_at_top (A-F)))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 233 | using tendsto_mono by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 234 | have "((\<lambda>G. sum f (F' \<union> G)) \<longlongrightarrow> x) (finite_subsets_at_top (A - F))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 235 | if "((sum f \<circ> (\<union>) F') \<longlongrightarrow> x) (finite_subsets_at_top (A - F))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 236 | using that unfolding o_def by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 237 | hence "((\<lambda>G. sum f (F' \<union> G)) \<longlongrightarrow> x) (finite_subsets_at_top (A-F))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 238 | using tendsto_compose_filtermap [symmetric] | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 239 | by (simp add: \<open>(sum f \<longlongrightarrow> x) (filtermap ((\<union>) F') (finite_subsets_at_top (A - F)))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 240 | tendsto_compose_filtermap) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 241 | have "\<forall>Y. finite Y \<and> Y \<subseteq> A - F \<longrightarrow> sum f (F' \<union> Y) = sum f F' + sum f Y" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 242 | by (metis Diff_disjoint Int_Diff \<open>A - F = A - F'\<close> \<open>finite F'\<close> inf.orderE sum.union_disjoint) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 243 | hence "\<forall>\<^sub>F x in finite_subsets_at_top (A - F). sum f (F' \<union> x) = sum f F' + sum f x" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 244 | unfolding eventually_finite_subsets_at_top | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 245 |     using exI [where x = "{}"]
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 246 |     by (simp add: \<open>\<And>P. P {} \<Longrightarrow> \<exists>x. P x\<close>) 
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 247 | hence "((\<lambda>G. sum f F' + sum f G) \<longlongrightarrow> x) (finite_subsets_at_top (A-F))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 248 | using tendsto_cong [THEN iffD1 , rotated] | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 249 | \<open>((\<lambda>G. sum f (F' \<union> G)) \<longlongrightarrow> x) (finite_subsets_at_top (A - F))\<close> by fastforce | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 250 | hence "((\<lambda>G. sum f F' + sum f G) \<longlongrightarrow> sum f F' + (x-sum f F')) (finite_subsets_at_top (A-F))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 251 | by simp | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 252 | hence "(sum f \<longlongrightarrow> x - sum f F') (finite_subsets_at_top (A-F))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 253 | using tendsto_add_const_iff by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 254 | thus "f summable_on (A - F)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 255 | unfolding summable_on_def has_sum_def by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 256 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 257 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 258 | lemma | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 259 |   fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add}"
 | 
| 77359 | 260 | assumes \<open>(f has_sum b) B\<close> and \<open>(f has_sum a) A\<close> and AB: "A \<subseteq> B" | 
| 261 | shows has_sum_Diff: "(f has_sum (b - a)) (B - A)" | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 262 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 263 | have finite_subsets1: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 264 | "finite_subsets_at_top (B - A) \<le> filtermap (\<lambda>F. F - A) (finite_subsets_at_top B)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 265 | proof (rule filter_leI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 266 | fix P assume "eventually P (filtermap (\<lambda>F. F - A) (finite_subsets_at_top B))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 267 | then obtain X where "finite X" and "X \<subseteq> B" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 268 | and P: "finite Y \<and> X \<subseteq> Y \<and> Y \<subseteq> B \<longrightarrow> P (Y - A)" for Y | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 269 | unfolding eventually_filtermap eventually_finite_subsets_at_top by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 270 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 271 | hence "finite (X-A)" and "X-A \<subseteq> B - A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 272 | by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 273 | moreover have "finite Y \<and> X-A \<subseteq> Y \<and> Y \<subseteq> B - A \<longrightarrow> P Y" for Y | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 274 | using P[where Y="Y\<union>X"] \<open>finite X\<close> \<open>X \<subseteq> B\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 275 | by (metis Diff_subset Int_Diff Un_Diff finite_Un inf.orderE le_sup_iff sup.orderE sup_ge2) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 276 | ultimately show "eventually P (finite_subsets_at_top (B - A))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 277 | unfolding eventually_finite_subsets_at_top by meson | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 278 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 279 | have finite_subsets2: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 280 | "filtermap (\<lambda>F. F \<inter> A) (finite_subsets_at_top B) \<le> finite_subsets_at_top A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 281 | apply (rule filter_leI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 282 | using assms unfolding eventually_filtermap eventually_finite_subsets_at_top | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 283 | by (metis Int_subset_iff finite_Int inf_le2 subset_trans) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 284 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 285 | from assms(1) have limB: "(sum f \<longlongrightarrow> b) (finite_subsets_at_top B)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 286 | using has_sum_def by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 287 | from assms(2) have limA: "(sum f \<longlongrightarrow> a) (finite_subsets_at_top A)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 288 | using has_sum_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 289 | have "((\<lambda>F. sum f (F\<inter>A)) \<longlongrightarrow> a) (finite_subsets_at_top B)" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 290 | proof (subst asm_rl [of "(\<lambda>F. sum f (F\<inter>A)) = sum f \<circ> (\<lambda>F. F\<inter>A)"]) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 291 | show "(\<lambda>F. sum f (F \<inter> A)) = sum f \<circ> (\<lambda>F. F \<inter> A)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 292 | unfolding o_def by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 293 | show "((sum f \<circ> (\<lambda>F. F \<inter> A)) \<longlongrightarrow> a) (finite_subsets_at_top B)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 294 | unfolding o_def | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 295 | using tendsto_compose_filtermap finite_subsets2 limA tendsto_mono | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 296 | \<open>(\<lambda>F. sum f (F \<inter> A)) = sum f \<circ> (\<lambda>F. F \<inter> A)\<close> by fastforce | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 297 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 298 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 299 | with limB have "((\<lambda>F. sum f F - sum f (F\<inter>A)) \<longlongrightarrow> b - a) (finite_subsets_at_top B)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 300 | using tendsto_diff by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 301 | have "sum f X - sum f (X \<inter> A) = sum f (X - A)" if "finite X" and "X \<subseteq> B" for X :: "'a set" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 302 | using that by (metis add_diff_cancel_left' sum.Int_Diff) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 303 | hence "\<forall>\<^sub>F x in finite_subsets_at_top B. sum f x - sum f (x \<inter> A) = sum f (x - A)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 304 | by (rule eventually_finite_subsets_at_top_weakI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 305 | hence "((\<lambda>F. sum f (F-A)) \<longlongrightarrow> b - a) (finite_subsets_at_top B)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 306 | using tendsto_cong [THEN iffD1 , rotated] | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 307 | \<open>((\<lambda>F. sum f F - sum f (F \<inter> A)) \<longlongrightarrow> b - a) (finite_subsets_at_top B)\<close> by fastforce | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 308 | hence "(sum f \<longlongrightarrow> b - a) (filtermap (\<lambda>F. F-A) (finite_subsets_at_top B))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 309 | by (subst tendsto_compose_filtermap[symmetric], simp add: o_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 310 | thus ?thesis | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 311 | using finite_subsets1 has_sum_def tendsto_mono by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 312 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 313 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 314 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 315 | lemma | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 316 |   fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 317 | assumes "f summable_on B" and "f summable_on A" and "A \<subseteq> B" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 318 | shows summable_on_Diff: "f summable_on (B-A)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 319 | by (meson assms summable_on_def has_sum_Diff) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 320 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 321 | lemma | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 322 |   fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add,t2_space}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 323 | assumes "f summable_on B" and "f summable_on A" and AB: "A \<subseteq> B" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 324 | shows infsum_Diff: "infsum f (B - A) = infsum f B - infsum f A" | 
| 74639 
f831b6e589dc
tuned proofs -- avoid z3, which is unavailable on arm64-linux;
 wenzelm parents: 
74475diff
changeset | 325 | by (metis AB assms has_sum_Diff infsumI summable_on_def) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 326 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 327 | lemma has_sum_mono_neutral: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 328 |   fixes f :: "'a\<Rightarrow>'b::{ordered_comm_monoid_add,linorder_topology}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 329 | (* Does this really require a linorder topology? (Instead of order topology.) *) | 
| 77359 | 330 | assumes \<open>(f has_sum a) A\<close> and "(g has_sum b) B" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 331 | assumes \<open>\<And>x. x \<in> A\<inter>B \<Longrightarrow> f x \<le> g x\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 332 | assumes \<open>\<And>x. x \<in> A-B \<Longrightarrow> f x \<le> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 333 | assumes \<open>\<And>x. x \<in> B-A \<Longrightarrow> g x \<ge> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 334 | shows "a \<le> b" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 335 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 336 | define f' g' where \<open>f' x = (if x \<in> A then f x else 0)\<close> and \<open>g' x = (if x \<in> B then g x else 0)\<close> for x | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 337 | have [simp]: \<open>f summable_on A\<close> \<open>g summable_on B\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 338 | using assms(1,2) summable_on_def by auto | 
| 77359 | 339 | have \<open>(f' has_sum a) (A\<union>B)\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 340 | by (smt (verit, best) DiffE IntE Un_iff f'_def assms(1) has_sum_cong_neutral) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 341 | then have f'_lim: \<open>(sum f' \<longlongrightarrow> a) (finite_subsets_at_top (A\<union>B))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 342 | by (meson has_sum_def) | 
| 77359 | 343 | have \<open>(g' has_sum b) (A\<union>B)\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 344 | by (smt (verit, best) DiffD1 DiffD2 IntE UnCI g'_def assms(2) has_sum_cong_neutral) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 345 | then have g'_lim: \<open>(sum g' \<longlongrightarrow> b) (finite_subsets_at_top (A\<union>B))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 346 | using has_sum_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 347 | |
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 348 | have "\<And>X i. \<lbrakk>X \<subseteq> A \<union> B; i \<in> X\<rbrakk> \<Longrightarrow> f' i \<le> g' i" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 349 | using assms by (auto simp: f'_def g'_def) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 350 | then have \<open>\<forall>\<^sub>F x in finite_subsets_at_top (A \<union> B). sum f' x \<le> sum g' x\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 351 | by (intro eventually_finite_subsets_at_top_weakI sum_mono) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 352 | then show ?thesis | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 353 | using f'_lim finite_subsets_at_top_neq_bot g'_lim tendsto_le by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 354 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 355 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 356 | lemma infsum_mono_neutral: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 357 |   fixes f :: "'a\<Rightarrow>'b::{ordered_comm_monoid_add,linorder_topology}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 358 | assumes "f summable_on A" and "g summable_on B" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 359 | assumes \<open>\<And>x. x \<in> A\<inter>B \<Longrightarrow> f x \<le> g x\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 360 | assumes \<open>\<And>x. x \<in> A-B \<Longrightarrow> f x \<le> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 361 | assumes \<open>\<And>x. x \<in> B-A \<Longrightarrow> g x \<ge> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 362 | shows "infsum f A \<le> infsum g B" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 363 | by (smt (verit, best) assms has_sum_infsum has_sum_mono_neutral) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 364 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 365 | lemma has_sum_mono: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 366 |   fixes f :: "'a\<Rightarrow>'b::{ordered_comm_monoid_add,linorder_topology}"
 | 
| 77359 | 367 | assumes "(f has_sum x) A" and "(g has_sum y) A" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 368 | assumes \<open>\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 369 | shows "x \<le> y" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 370 | using assms has_sum_mono_neutral by force | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 371 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 372 | lemma infsum_mono: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 373 |   fixes f :: "'a\<Rightarrow>'b::{ordered_comm_monoid_add,linorder_topology}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 374 | assumes "f summable_on A" and "g summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 375 | assumes \<open>\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 376 | shows "infsum f A \<le> infsum g A" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 377 | by (meson assms has_sum_infsum has_sum_mono) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 378 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 379 | lemma has_sum_finite[simp]: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 380 | assumes "finite F" | 
| 77359 | 381 | shows "(f has_sum (sum f F)) F" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 382 | using assms | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 383 | by (auto intro: tendsto_Lim simp: finite_subsets_at_top_finite infsum_def has_sum_def principal_eq_bot_iff) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 384 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 385 | lemma summable_on_finite[simp]: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 386 |   fixes f :: \<open>'a \<Rightarrow> 'b::{comm_monoid_add,topological_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 387 | assumes "finite F" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 388 | shows "f summable_on F" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 389 | using assms summable_on_def has_sum_finite by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 390 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 391 | lemma infsum_finite[simp]: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 392 | assumes "finite F" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 393 | shows "infsum f F = sum f F" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 394 | by (simp add: assms infsumI) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 395 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 396 | lemma has_sum_finite_approximation: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 397 |   fixes f :: "'a \<Rightarrow> 'b::{comm_monoid_add,metric_space}"
 | 
| 77359 | 398 | assumes "(f has_sum x) A" and "\<epsilon> > 0" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 399 | shows "\<exists>F. finite F \<and> F \<subseteq> A \<and> dist (sum f F) x \<le> \<epsilon>" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 400 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 401 | have "(sum f \<longlongrightarrow> x) (finite_subsets_at_top A)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 402 | by (meson assms(1) has_sum_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 403 | hence *: "\<forall>\<^sub>F F in (finite_subsets_at_top A). dist (sum f F) x < \<epsilon>" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 404 | using assms(2) by (rule tendstoD) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 405 | thus ?thesis | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 406 | unfolding eventually_finite_subsets_at_top by fastforce | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 407 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 408 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 409 | lemma infsum_finite_approximation: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 410 |   fixes f :: "'a \<Rightarrow> 'b::{comm_monoid_add,metric_space}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 411 | assumes "f summable_on A" and "\<epsilon> > 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 412 | shows "\<exists>F. finite F \<and> F \<subseteq> A \<and> dist (sum f F) (infsum f A) \<le> \<epsilon>" | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 413 | proof - | 
| 77359 | 414 | from assms have "(f has_sum (infsum f A)) A" | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 415 | by (simp add: summable_iff_has_sum_infsum) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 416 | from this and \<open>\<epsilon> > 0\<close> show ?thesis | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 417 | by (rule has_sum_finite_approximation) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 418 | qed | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 419 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 420 | lemma abs_summable_summable: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 421 | fixes f :: \<open>'a \<Rightarrow> 'b :: banach\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 422 | assumes \<open>f abs_summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 423 | shows \<open>f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 424 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 425 | from assms obtain L where lim: \<open>(sum (\<lambda>x. norm (f x)) \<longlongrightarrow> L) (finite_subsets_at_top A)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 426 | unfolding has_sum_def summable_on_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 427 | then have *: \<open>cauchy_filter (filtermap (sum (\<lambda>x. norm (f x))) (finite_subsets_at_top A))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 428 | by (auto intro!: nhds_imp_cauchy_filter simp: filterlim_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 429 | have \<open>\<exists>P. eventually P (finite_subsets_at_top A) \<and> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 430 | (\<forall>F F'. P F \<and> P F' \<longrightarrow> dist (sum f F) (sum f F') < e)\<close> if \<open>e>0\<close> for e | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 431 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 432 | define d P where \<open>d = e/4\<close> and \<open>P F \<longleftrightarrow> finite F \<and> F \<subseteq> A \<and> dist (sum (\<lambda>x. norm (f x)) F) L < d\<close> for F | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 433 | then have \<open>d > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 434 | by (simp add: d_def that) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 435 | have ev_P: \<open>eventually P (finite_subsets_at_top A)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 436 | using lim | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 437 | by (auto simp add: P_def[abs_def] \<open>0 < d\<close> eventually_conj_iff eventually_finite_subsets_at_top_weakI tendsto_iff) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 438 | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 439 | moreover have \<open>dist (sum f F1) (sum f F2) < e\<close> if \<open>P F1\<close> and \<open>P F2\<close> for F1 F2 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 440 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 441 | from ev_P | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 442 | obtain F' where \<open>finite F'\<close> and \<open>F' \<subseteq> A\<close> and P_sup_F': \<open>finite F \<and> F \<supseteq> F' \<and> F \<subseteq> A \<Longrightarrow> P F\<close> for F | 
| 74639 
f831b6e589dc
tuned proofs -- avoid z3, which is unavailable on arm64-linux;
 wenzelm parents: 
74475diff
changeset | 443 | by atomize_elim (simp add: eventually_finite_subsets_at_top) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 444 | define F where \<open>F = F' \<union> F1 \<union> F2\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 445 | have \<open>finite F\<close> and \<open>F \<subseteq> A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 446 | using F_def P_def[abs_def] that \<open>finite F'\<close> \<open>F' \<subseteq> A\<close> by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 447 | have dist_F: \<open>dist (sum (\<lambda>x. norm (f x)) F) L < d\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 448 | by (metis F_def \<open>F \<subseteq> A\<close> P_def P_sup_F' \<open>finite F\<close> le_supE order_refl) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 449 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 450 | have dist_F_subset: \<open>dist (sum f F) (sum f F') < 2*d\<close> if F': \<open>F' \<subseteq> F\<close> \<open>P F'\<close> for F' | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 451 | proof - | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 452 | have \<open>dist (sum f F) (sum f F') = norm (sum f (F-F'))\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 453 | unfolding dist_norm using \<open>finite F\<close> F' by (subst sum_diff) auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 454 | also have \<open>\<dots> \<le> norm (\<Sum>x\<in>F-F'. norm (f x))\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 455 | by (rule order.trans[OF sum_norm_le[OF order.refl]]) auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 456 | also have \<open>\<dots> = dist (\<Sum>x\<in>F. norm (f x)) (\<Sum>x\<in>F'. norm (f x))\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 457 | unfolding dist_norm using \<open>finite F\<close> F' by (subst sum_diff) auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 458 | also have \<open>\<dots> < 2 * d\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 459 | using dist_F F' unfolding P_def dist_norm real_norm_def by linarith | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 460 | finally show \<open>dist (sum f F) (sum f F') < 2*d\<close> . | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 461 | qed | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 462 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 463 | have \<open>dist (sum f F1) (sum f F2) \<le> dist (sum f F) (sum f F1) + dist (sum f F) (sum f F2)\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 464 | by (rule dist_triangle3) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 465 | also have \<open>\<dots> < 2 * d + 2 * d\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 466 | by (intro add_strict_mono dist_F_subset that) (auto simp: F_def) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 467 | also have \<open>\<dots> \<le> e\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 468 | by (auto simp: d_def) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 469 | finally show \<open>dist (sum f F1) (sum f F2) < e\<close> . | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 470 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 471 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 472 | using ev_P by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 473 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 474 | then have \<open>cauchy_filter (filtermap (sum f) (finite_subsets_at_top A))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 475 | by (simp add: cauchy_filter_metric_filtermap) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 476 | moreover have "complete (UNIV::'b set)" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 477 | by (meson Cauchy_convergent UNIV_I complete_def convergent_def) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 478 | ultimately obtain L' where \<open>(sum f \<longlongrightarrow> L') (finite_subsets_at_top A)\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 479 | using complete_uniform[where S=UNIV] by (force simp add: filterlim_def) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 480 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 481 | using summable_on_def has_sum_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 482 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 483 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 484 | text \<open>The converse of @{thm [source] abs_summable_summable} does not hold:
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 485 | Consider the Hilbert space of square-summable sequences. | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 486 | Let $e_i$ denote the sequence with 1 in the $i$th position and 0 elsewhere. | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 487 | Let $f(i) := e_i/i$ for $i\geq1$. We have \<^term>\<open>\<not> f abs_summable_on UNIV\<close> because $\lVert f(i)\rVert=1/i$ | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 488 | and thus the sum over $\lVert f(i)\rVert$ diverges. On the other hand, we have \<^term>\<open>f summable_on UNIV\<close>; | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 489 | the limit is the sequence with $1/i$ in the $i$th position. | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 490 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 491 | (We have not formalized this separating example here because to the best of our knowledge, | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 492 | this Hilbert space has not been formalized in Isabelle/HOL yet.)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 493 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 494 | lemma norm_has_sum_bound: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 495 | fixes f :: "'b \<Rightarrow> 'a::real_normed_vector" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 496 | and A :: "'b set" | 
| 77359 | 497 | assumes "((\<lambda>x. norm (f x)) has_sum n) A" | 
| 498 | assumes "(f has_sum a) A" | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 499 | shows "norm a \<le> n" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 500 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 501 | have "norm a \<le> n + \<epsilon>" if "\<epsilon>>0" for \<epsilon> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 502 | proof- | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 503 | have "\<exists>F. norm (a - sum f F) \<le> \<epsilon> \<and> finite F \<and> F \<subseteq> A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 504 | using has_sum_finite_approximation[where A=A and f=f and \<epsilon>="\<epsilon>"] assms \<open>0 < \<epsilon>\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 505 | by (metis dist_commute dist_norm) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 506 | then obtain F where "norm (a - sum f F) \<le> \<epsilon>" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 507 | and "finite F" and "F \<subseteq> A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 508 | by (simp add: atomize_elim) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 509 | hence "norm a \<le> norm (sum f F) + \<epsilon>" | 
| 74642 
8af77cb50c6d
tuned proofs -- avoid z3, which is unavailable on arm64-linux;
 wenzelm parents: 
74639diff
changeset | 510 | by (metis add.commute diff_add_cancel dual_order.refl norm_triangle_mono) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 511 | also have "\<dots> \<le> sum (\<lambda>x. norm (f x)) F + \<epsilon>" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 512 | using norm_sum by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 513 | also have "\<dots> \<le> n + \<epsilon>" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 514 | proof (intro add_right_mono [OF has_sum_mono_neutral]) | 
| 77359 | 515 | show "((\<lambda>x. norm (f x)) has_sum (\<Sum>x\<in>F. norm (f x))) F" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 516 | by (simp add: \<open>finite F\<close>) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 517 | qed (use \<open>F \<subseteq> A\<close> assms in auto) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 518 | finally show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 519 | by assumption | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 520 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 521 | thus ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 522 | using linordered_field_class.field_le_epsilon by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 523 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 524 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 525 | lemma norm_infsum_bound: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 526 | fixes f :: "'b \<Rightarrow> 'a::real_normed_vector" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 527 | and A :: "'b set" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 528 | assumes "f abs_summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 529 | shows "norm (infsum f A) \<le> infsum (\<lambda>x. norm (f x)) A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 530 | proof (cases "f summable_on A") | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 531 | case True | 
| 77359 | 532 | have "((\<lambda>x. norm (f x)) has_sum (\<Sum>\<^sub>\<infinity>x\<in>A. norm (f x))) A" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 533 | by (simp add: assms) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 534 | then show ?thesis | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 535 | by (metis True has_sum_infsum norm_has_sum_bound) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 536 | next | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 537 | case False | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 538 | obtain t where t_def: "(sum (\<lambda>x. norm (f x)) \<longlongrightarrow> t) (finite_subsets_at_top A)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 539 | using assms unfolding summable_on_def has_sum_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 540 | have sumpos: "sum (\<lambda>x. norm (f x)) X \<ge> 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 541 | for X | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 542 | by (simp add: sum_nonneg) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 543 | have tgeq0:"t \<ge> 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 544 | proof(rule ccontr) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 545 |     define S::"real set" where "S = {s. s < 0}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 546 | assume "\<not> 0 \<le> t" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 547 | hence "t < 0" by simp | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 548 | hence "t \<in> S" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 549 | unfolding S_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 550 | moreover have "open S" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 551 | by (metis S_def lessThan_def open_real_lessThan) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 552 | ultimately have "\<forall>\<^sub>F X in finite_subsets_at_top A. (\<Sum>x\<in>X. norm (f x)) \<in> S" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 553 | using t_def unfolding tendsto_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 554 | hence "\<exists>X. (\<Sum>x\<in>X. norm (f x)) \<in> S" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 555 | by (metis (no_types, lifting) eventually_mono filterlim_iff finite_subsets_at_top_neq_bot tendsto_Lim) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 556 | then obtain X where "(\<Sum>x\<in>X. norm (f x)) \<in> S" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 557 | by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 558 | hence "(\<Sum>x\<in>X. norm (f x)) < 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 559 | unfolding S_def by auto | 
| 74642 
8af77cb50c6d
tuned proofs -- avoid z3, which is unavailable on arm64-linux;
 wenzelm parents: 
74639diff
changeset | 560 | thus False by (simp add: leD sumpos) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 561 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 562 | have "\<exists>!h. (sum (\<lambda>x. norm (f x)) \<longlongrightarrow> h) (finite_subsets_at_top A)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 563 | using t_def finite_subsets_at_top_neq_bot tendsto_unique by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 564 | hence "t = (Topological_Spaces.Lim (finite_subsets_at_top A) (sum (\<lambda>x. norm (f x))))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 565 | using t_def unfolding Topological_Spaces.Lim_def | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 566 | by (metis the_equality) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 567 | hence "Lim (finite_subsets_at_top A) (sum (\<lambda>x. norm (f x))) \<ge> 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 568 | using tgeq0 by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 569 | thus ?thesis unfolding infsum_def | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 570 | using False by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 571 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 572 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 573 | lemma infsum_tendsto: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 574 | assumes \<open>f summable_on S\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 575 | shows \<open>((\<lambda>F. sum f F) \<longlongrightarrow> infsum f S) (finite_subsets_at_top S)\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 576 | using assms has_sum_def has_sum_infsum by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 577 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 578 | lemma has_sum_0: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 579 | assumes \<open>\<And>x. x\<in>M \<Longrightarrow> f x = 0\<close> | 
| 77359 | 580 | shows \<open>(f has_sum 0) M\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 581 | by (metis assms finite.intros(1) has_sum_cong has_sum_cong_neutral has_sum_finite sum.neutral_const) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 582 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 583 | lemma summable_on_0: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 584 | assumes \<open>\<And>x. x\<in>M \<Longrightarrow> f x = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 585 | shows \<open>f summable_on M\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 586 | using assms summable_on_def has_sum_0 by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 587 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 588 | lemma infsum_0: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 589 | assumes \<open>\<And>x. x\<in>M \<Longrightarrow> f x = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 590 | shows \<open>infsum f M = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 591 | by (metis assms finite_subsets_at_top_neq_bot infsum_def has_sum_0 has_sum_def tendsto_Lim) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 592 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 593 | text \<open>Variants of @{thm [source] infsum_0} etc. suitable as simp-rules\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 594 | lemma infsum_0_simp[simp]: \<open>infsum (\<lambda>_. 0) M = 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 595 | by (simp_all add: infsum_0) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 596 | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 597 | lemma summable_on_0_simp[simp]: \<open>(\<lambda>_. 0) summable_on M\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 598 | by (simp_all add: summable_on_0) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 599 | |
| 77359 | 600 | lemma has_sum_0_simp[simp]: \<open>((\<lambda>_. 0) has_sum 0) M\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 601 | by (simp_all add: has_sum_0) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 602 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 603 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 604 | lemma has_sum_add: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 605 |   fixes f g :: "'a \<Rightarrow> 'b::{topological_comm_monoid_add}"
 | 
| 77359 | 606 | assumes \<open>(f has_sum a) A\<close> | 
| 607 | assumes \<open>(g has_sum b) A\<close> | |
| 608 | shows \<open>((\<lambda>x. f x + g x) has_sum (a + b)) A\<close> | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 609 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 610 | from assms have lim_f: \<open>(sum f \<longlongrightarrow> a) (finite_subsets_at_top A)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 611 | and lim_g: \<open>(sum g \<longlongrightarrow> b) (finite_subsets_at_top A)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 612 | by (simp_all add: has_sum_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 613 | then have lim: \<open>(sum (\<lambda>x. f x + g x) \<longlongrightarrow> a + b) (finite_subsets_at_top A)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 614 | unfolding sum.distrib by (rule tendsto_add) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 615 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 616 | by (simp_all add: has_sum_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 617 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 618 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 619 | lemma summable_on_add: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 620 |   fixes f g :: "'a \<Rightarrow> 'b::{topological_comm_monoid_add}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 621 | assumes \<open>f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 622 | assumes \<open>g summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 623 | shows \<open>(\<lambda>x. f x + g x) summable_on A\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 624 | by (metis (full_types) assms summable_on_def has_sum_add) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 625 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 626 | lemma infsum_add: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 627 |   fixes f g :: "'a \<Rightarrow> 'b::{topological_comm_monoid_add, t2_space}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 628 | assumes \<open>f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 629 | assumes \<open>g summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 630 | shows \<open>infsum (\<lambda>x. f x + g x) A = infsum f A + infsum g A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 631 | proof - | 
| 77359 | 632 | have \<open>((\<lambda>x. f x + g x) has_sum (infsum f A + infsum g A)) A\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 633 | by (simp add: assms has_sum_add) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 634 | then show ?thesis | 
| 74639 
f831b6e589dc
tuned proofs -- avoid z3, which is unavailable on arm64-linux;
 wenzelm parents: 
74475diff
changeset | 635 | using infsumI by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 636 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 637 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 638 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 639 | lemma has_sum_Un_disjoint: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 640 | fixes f :: "'a \<Rightarrow> 'b::topological_comm_monoid_add" | 
| 77359 | 641 | assumes "(f has_sum a) A" | 
| 642 | assumes "(f has_sum b) B" | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 643 |   assumes disj: "A \<inter> B = {}"
 | 
| 77359 | 644 | shows \<open>(f has_sum (a + b)) (A \<union> B)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 645 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 646 | define fA fB where \<open>fA x = (if x \<in> A then f x else 0)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 647 | and \<open>fB x = (if x \<notin> A then f x else 0)\<close> for x | 
| 77359 | 648 | have fA: \<open>(fA has_sum a) (A \<union> B)\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 649 | by (smt (verit, ccfv_SIG) DiffD1 DiffD2 UnCI fA_def assms(1) has_sum_cong_neutral inf_sup_absorb) | 
| 77359 | 650 | have fB: \<open>(fB has_sum b) (A \<union> B)\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 651 | by (smt (verit, best) DiffD1 DiffD2 IntE Un_iff fB_def assms(2) disj disjoint_iff has_sum_cong_neutral) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 652 | have fAB: \<open>f x = fA x + fB x\<close> for x | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 653 | unfolding fA_def fB_def by simp | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 654 | show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 655 | unfolding fAB | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 656 | using fA fB by (rule has_sum_add) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 657 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 658 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 659 | lemma summable_on_Un_disjoint: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 660 | fixes f :: "'a \<Rightarrow> 'b::topological_comm_monoid_add" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 661 | assumes "f summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 662 | assumes "f summable_on B" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 663 |   assumes disj: "A \<inter> B = {}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 664 | shows \<open>f summable_on (A \<union> B)\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 665 | by (meson assms disj summable_on_def has_sum_Un_disjoint) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 666 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 667 | lemma infsum_Un_disjoint: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 668 |   fixes f :: "'a \<Rightarrow> 'b::{topological_comm_monoid_add, t2_space}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 669 | assumes "f summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 670 | assumes "f summable_on B" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 671 |   assumes disj: "A \<inter> B = {}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 672 | shows \<open>infsum f (A \<union> B) = infsum f A + infsum f B\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 673 | by (intro infsumI has_sum_Un_disjoint has_sum_infsum assms) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 674 | |
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 675 | lemma norm_summable_imp_has_sum: | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 676 | fixes f :: "nat \<Rightarrow> 'a :: banach" | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 677 | assumes "summable (\<lambda>n. norm (f n))" and "f sums S" | 
| 77359 | 678 | shows "(f has_sum S) (UNIV :: nat set)" | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 679 | unfolding has_sum_def tendsto_iff eventually_finite_subsets_at_top | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 680 | proof clarsimp | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 681 | fix \<epsilon>::real | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 682 | assume "\<epsilon> > 0" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 683 | from assms obtain S' where S': "(\<lambda>n. norm (f n)) sums S'" | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 684 | by (auto simp: summable_def) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 685 | with \<open>\<epsilon> > 0\<close> obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> \<bar>S' - (\<Sum>i<n. norm (f i))\<bar> < \<epsilon>" | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 686 | by (auto simp: tendsto_iff eventually_at_top_linorder sums_def dist_norm abs_minus_commute) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 687 |   have "dist (sum f Y) S < \<epsilon>" if "finite Y" "{..<N} \<subseteq> Y" for Y
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 688 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 689 | from that have "(\<lambda>n. if n \<in> Y then 0 else f n) sums (S - sum f Y)" | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 690 | by (intro sums_If_finite_set'[OF \<open>f sums S\<close>]) (auto simp: sum_negf) | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 691 | hence "S - sum f Y = (\<Sum>n. if n \<in> Y then 0 else f n)" | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 692 | by (simp add: sums_iff) | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 693 | also have "norm \<dots> \<le> (\<Sum>n. norm (if n \<in> Y then 0 else f n))" | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 694 | by (rule summable_norm[OF summable_comparison_test'[OF assms(1)]]) auto | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 695 | also have "\<dots> \<le> (\<Sum>n. if n < N then 0 else norm (f n))" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 696 | using that by (intro suminf_le summable_comparison_test'[OF assms(1)]) auto | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 697 |     also have "(\<lambda>n. if n \<in> {..<N} then 0 else norm (f n)) sums (S' - (\<Sum>i<N. norm (f i)))" 
 | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 698 | by (intro sums_If_finite_set'[OF S']) (auto simp: sum_negf) | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 699 | hence "(\<Sum>n. if n < N then 0 else norm (f n)) = S' - (\<Sum>i<N. norm (f i))" | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 700 | by (simp add: sums_iff) | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 701 | also have "S' - (\<Sum>i<N. norm (f i)) \<le> \<bar>S' - (\<Sum>i<N. norm (f i))\<bar>" by simp | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 702 | also have "\<dots> < \<epsilon>" by (rule N) auto | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 703 | finally show ?thesis by (simp add: dist_norm norm_minus_commute) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 704 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 705 | then show "\<exists>X. finite X \<and> (\<forall>Y. finite Y \<and> X \<subseteq> Y \<longrightarrow> dist (sum f Y) S < \<epsilon>)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 706 | by (meson finite_lessThan subset_UNIV) | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 707 | qed | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 708 | |
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 709 | lemma norm_summable_imp_summable_on: | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 710 | fixes f :: "nat \<Rightarrow> 'a :: banach" | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 711 | assumes "summable (\<lambda>n. norm (f n))" | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 712 | shows "f summable_on UNIV" | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 713 | using norm_summable_imp_has_sum[OF assms, of "suminf f"] assms | 
| 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74791diff
changeset | 714 | by (auto simp: sums_iff summable_on_def dest: summable_norm_cancel) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 715 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 716 | text \<open>The following lemma indeed needs a complete space (as formalized by the premise \<^term>\<open>complete UNIV\<close>). | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 717 | The following two counterexamples show this: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 718 |   \begin{itemize}
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 719 | \item Consider the real vector space $V$ of sequences with finite support, and with the $\ell_2$-norm (sum of squares). | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 720 | Let $e_i$ denote the sequence with a $1$ at position $i$. | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 721 |       Let $f : \mathbb Z \to V$ be defined as $f(n) := e_{\lvert n\rvert} / n$ (with $f(0) := 0$).
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 722 |       We have that $\sum_{n\in\mathbb Z} f(n) = 0$ (it even converges absolutely). 
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 723 |       But $\sum_{n\in\mathbb N} f(n)$ does not exist (it would converge against a sequence with infinite support).
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 724 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 725 |   \item Let $f$ be a positive rational valued function such that $\sum_{x\in B} f(x)$ is $\sqrt 2$ and $\sum_{x\in A} f(x)$ is 1 (over the reals, with $A\subseteq B$).
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 726 |       Then $\sum_{x\in B} f(x)$ does not exist over the rationals. But $\sum_{x\in A} f(x)$ exists.
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 727 |   \end{itemize}
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 728 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 729 | The lemma also requires uniform continuity of the addition. And example of a topological group with continuous | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 730 | but not uniformly continuous addition would be the positive reals with the usual multiplication as the addition. | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 731 | We do not know whether the lemma would also hold for such topological groups.\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 732 | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 733 | lemma summable_on_subset_aux: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 734 |   fixes A B and f :: \<open>'a \<Rightarrow> 'b::{ab_group_add, uniform_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 735 | assumes \<open>complete (UNIV :: 'b set)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 736 | assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'b,y). x+y)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 737 | assumes \<open>f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 738 | assumes \<open>B \<subseteq> A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 739 | shows \<open>f summable_on B\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 740 | proof - | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 741 | let ?filter_fB = \<open>filtermap (sum f) (finite_subsets_at_top B)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 742 | from \<open>f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 743 | obtain S where \<open>(sum f \<longlongrightarrow> S) (finite_subsets_at_top A)\<close> (is \<open>(sum f \<longlongrightarrow> S) ?filter_A\<close>) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 744 | using summable_on_def has_sum_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 745 | then have cauchy_fA: \<open>cauchy_filter (filtermap (sum f) (finite_subsets_at_top A))\<close> (is \<open>cauchy_filter ?filter_fA\<close>) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 746 | by (auto intro!: nhds_imp_cauchy_filter simp: filterlim_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 747 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 748 | have \<open>cauchy_filter (filtermap (sum f) (finite_subsets_at_top B))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 749 | proof (unfold cauchy_filter_def, rule filter_leI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 750 |     fix E :: \<open>('b\<times>'b) \<Rightarrow> bool\<close> assume \<open>eventually E uniformity\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 751 | then obtain E' where \<open>eventually E' uniformity\<close> and E'E'E: \<open>E' (x, y) \<longrightarrow> E' (y, z) \<longrightarrow> E (x, z)\<close> for x y z | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 752 | using uniformity_trans by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 753 | obtain D where \<open>eventually D uniformity\<close> and DE: \<open>D (x, y) \<Longrightarrow> E' (x+c, y+c)\<close> for x y c | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 754 | using plus_cont \<open>eventually E' uniformity\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 755 | unfolding uniformly_continuous_on_uniformity filterlim_def le_filter_def uniformity_prod_def | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 756 | by (auto simp: case_prod_beta eventually_filtermap eventually_prod_same uniformity_refl) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 757 | have DE': "E' (x, y)" if "D (x + c, y + c)" for x y c | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 758 | using DE[of "x + c" "y + c" "-c"] that by simp | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 759 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 760 | from \<open>eventually D uniformity\<close> and cauchy_fA have \<open>eventually D (?filter_fA \<times>\<^sub>F ?filter_fA)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 761 | unfolding cauchy_filter_def le_filter_def by simp | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 762 | then obtain P1 P2 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 763 | where ev_P1: \<open>eventually (\<lambda>F. P1 (sum f F)) ?filter_A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 764 | and ev_P2: \<open>eventually (\<lambda>F. P2 (sum f F)) ?filter_A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 765 | and P1P2E: \<open>P1 x \<Longrightarrow> P2 y \<Longrightarrow> D (x, y)\<close> for x y | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 766 | unfolding eventually_prod_filter eventually_filtermap | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 767 | by auto | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 768 | from ev_P1 obtain F1 where F1: \<open>finite F1\<close> \<open>F1 \<subseteq> A\<close> \<open>\<And>F. F\<supseteq>F1 \<Longrightarrow> finite F \<Longrightarrow> F\<subseteq>A \<Longrightarrow> P1 (sum f F)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 769 | by (metis eventually_finite_subsets_at_top) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 770 | from ev_P2 obtain F2 where F2: \<open>finite F2\<close> \<open>F2 \<subseteq> A\<close> \<open>\<And>F. F\<supseteq>F2 \<Longrightarrow> finite F \<Longrightarrow> F\<subseteq>A \<Longrightarrow> P2 (sum f F)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 771 | by (metis eventually_finite_subsets_at_top) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 772 | define F0 F0A F0B where \<open>F0 \<equiv> F1 \<union> F2\<close> and \<open>F0A \<equiv> F0 - B\<close> and \<open>F0B \<equiv> F0 \<inter> B\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 773 | have [simp]: \<open>finite F0\<close> \<open>F0 \<subseteq> A\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 774 | using \<open>F1 \<subseteq> A\<close> \<open>F2 \<subseteq> A\<close> \<open>finite F1\<close> \<open>finite F2\<close> unfolding F0_def by blast+ | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 775 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 776 | have *: "E' (sum f F1', sum f F2')" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 777 | if "F1'\<supseteq>F0B" "F2'\<supseteq>F0B" "finite F1'" "finite F2'" "F1'\<subseteq>B" "F2'\<subseteq>B" for F1' F2' | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 778 | proof (intro DE'[where c = "sum f F0A"] P1P2E) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 779 | have "P1 (sum f (F1' \<union> F0A))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 780 | using that assms F1(1,2) F2(1,2) by (intro F1) (auto simp: F0A_def F0B_def F0_def) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 781 | thus "P1 (sum f F1' + sum f F0A)" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 782 | by (subst (asm) sum.union_disjoint) (use that in \<open>auto simp: F0A_def\<close>) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 783 | next | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 784 | have "P2 (sum f (F2' \<union> F0A))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 785 | using that assms F1(1,2) F2(1,2) by (intro F2) (auto simp: F0A_def F0B_def F0_def) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 786 | thus "P2 (sum f F2' + sum f F0A)" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 787 | by (subst (asm) sum.union_disjoint) (use that in \<open>auto simp: F0A_def\<close>) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 788 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 789 | |
| 77388 | 790 | have "eventually (\<lambda>x. E' (x, sum f F0B)) (filtermap (sum f) (finite_subsets_at_top B))" | 
| 791 | and "eventually (\<lambda>x. E' (sum f F0B, x)) (filtermap (sum f) (finite_subsets_at_top B))" | |
| 792 | unfolding eventually_filtermap eventually_finite_subsets_at_top | |
| 793 | by (rule exI[of _ F0B]; use * in \<open>force simp: F0B_def\<close>)+ | |
| 794 | then | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 795 | show \<open>eventually E (?filter_fB \<times>\<^sub>F ?filter_fB)\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 796 | unfolding eventually_prod_filter | 
| 77388 | 797 | using E'E'E by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 798 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 799 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 800 | then obtain x where \<open>?filter_fB \<le> nhds x\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 801 | using cauchy_filter_complete_converges[of ?filter_fB UNIV] \<open>complete (UNIV :: _)\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 802 | by (auto simp: filtermap_bot_iff) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 803 | then have \<open>(sum f \<longlongrightarrow> x) (finite_subsets_at_top B)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 804 | by (auto simp: filterlim_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 805 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 806 | by (auto simp: summable_on_def has_sum_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 807 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 808 | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 809 | text \<open>A special case of @{thm [source] summable_on_subset_aux} for Banach spaces with fewer premises.\<close>
 | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 810 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 811 | lemma summable_on_subset_banach: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 812 | fixes A B and f :: \<open>'a \<Rightarrow> 'b::banach\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 813 | assumes \<open>f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 814 | assumes \<open>B \<subseteq> A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 815 | shows \<open>f summable_on B\<close> | 
| 77388 | 816 | by (meson Cauchy_convergent UNIV_I assms complete_def convergent_def isUCont_plus summable_on_subset_aux) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 817 | |
| 77359 | 818 | lemma has_sum_empty[simp]: \<open>(f has_sum 0) {}\<close>
 | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 819 | by (meson ex_in_conv has_sum_0) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 820 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 821 | lemma summable_on_empty[simp]: \<open>f summable_on {}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 822 | by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 823 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 824 | lemma infsum_empty[simp]: \<open>infsum f {} = 0\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 825 | by simp | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 826 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 827 | lemma sum_has_sum: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 828 | fixes f :: "'a \<Rightarrow> 'b::topological_comm_monoid_add" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 829 | assumes \<open>finite A\<close> | 
| 77359 | 830 | assumes \<open>\<And>a. a \<in> A \<Longrightarrow> (f has_sum (s a)) (B a)\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 831 |   assumes \<open>\<And>a a'. a\<in>A \<Longrightarrow> a'\<in>A \<Longrightarrow> a\<noteq>a' \<Longrightarrow> B a \<inter> B a' = {}\<close>
 | 
| 77359 | 832 | shows \<open>(f has_sum (sum s A)) (\<Union>a\<in>A. B a)\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 833 | using assms | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 834 | proof (induction) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 835 | case empty | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 836 | then show ?case | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 837 | by simp | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 838 | next | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 839 | case (insert x A) | 
| 77359 | 840 | have \<open>(f has_sum (s x)) (B x)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 841 | by (simp add: insert.prems) | 
| 77359 | 842 | moreover have IH: \<open>(f has_sum (sum s A)) (\<Union>a\<in>A. B a)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 843 | using insert by simp | 
| 77359 | 844 | ultimately have \<open>(f has_sum (s x + sum s A)) (B x \<union> (\<Union>a\<in>A. B a))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 845 | using insert by (intro has_sum_Un_disjoint) auto | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 846 | then show ?case | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 847 | using insert.hyps by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 848 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 849 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 850 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 851 | lemma summable_on_finite_union_disjoint: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 852 | fixes f :: "'a \<Rightarrow> 'b::topological_comm_monoid_add" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 853 | assumes finite: \<open>finite A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 854 | assumes conv: \<open>\<And>a. a \<in> A \<Longrightarrow> f summable_on (B a)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 855 |   assumes disj: \<open>\<And>a a'. a\<in>A \<Longrightarrow> a'\<in>A \<Longrightarrow> a\<noteq>a' \<Longrightarrow> B a \<inter> B a' = {}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 856 | shows \<open>f summable_on (\<Union>a\<in>A. B a)\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 857 | using sum_has_sum [of A f B] assms unfolding summable_on_def by metis | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 858 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 859 | lemma sum_infsum: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 860 |   fixes f :: "'a \<Rightarrow> 'b::{topological_comm_monoid_add, t2_space}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 861 | assumes finite: \<open>finite A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 862 | assumes conv: \<open>\<And>a. a \<in> A \<Longrightarrow> f summable_on (B a)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 863 |   assumes disj: \<open>\<And>a a'. a\<in>A \<Longrightarrow> a'\<in>A \<Longrightarrow> a\<noteq>a' \<Longrightarrow> B a \<inter> B a' = {}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 864 | shows \<open>sum (\<lambda>a. infsum f (B a)) A = infsum f (\<Union>a\<in>A. B a)\<close> | 
| 77388 | 865 | by (metis (no_types, lifting) assms has_sum_infsum infsumI sum_has_sum) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 866 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 867 | text \<open>The lemmas \<open>infsum_comm_additive_general\<close> and \<open>infsum_comm_additive\<close> (and variants) below both state that the infinite sum commutes with | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 868 | a continuous additive function. \<open>infsum_comm_additive_general\<close> is stated more for more general type classes | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 869 | at the expense of a somewhat less compact formulation of the premises. | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 870 | E.g., by avoiding the constant \<^const>\<open>additive\<close> which introduces an additional sort constraint | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 871 | (group instead of monoid). For example, extended reals (\<^typ>\<open>ereal\<close>, \<^typ>\<open>ennreal\<close>) are not covered | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 872 | by \<open>infsum_comm_additive\<close>.\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 873 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 874 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 875 | lemma has_sum_comm_additive_general: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 876 |   fixes f :: \<open>'b :: {comm_monoid_add,topological_space} \<Rightarrow> 'c :: {comm_monoid_add,topological_space}\<close>
 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 877 | assumes f_sum: \<open>\<And>F. finite F \<Longrightarrow> F \<subseteq> S \<Longrightarrow> sum (f \<circ> g) F = f (sum g F)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 878 | \<comment> \<open>Not using \<^const>\<open>additive\<close> because it would add sort constraint \<^class>\<open>ab_group_add\<close>\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 879 | assumes cont: \<open>f \<midarrow>x\<rightarrow> f x\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 880 |     \<comment> \<open>For \<^class>\<open>t2_space\<close>, this is equivalent to \<open>isCont f x\<close> by @{thm [source] isCont_def}.\<close>
 | 
| 77359 | 881 | assumes infsum: \<open>(g has_sum x) S\<close> | 
| 882 | shows \<open>((f \<circ> g) has_sum (f x)) S\<close> | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 883 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 884 | have \<open>(sum g \<longlongrightarrow> x) (finite_subsets_at_top S)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 885 | using infsum has_sum_def by blast | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 886 | then have \<open>((f \<circ> sum g) \<longlongrightarrow> f x) (finite_subsets_at_top S)\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 887 | by (meson cont filterlim_def tendsto_at_iff_tendsto_nhds tendsto_compose_filtermap tendsto_mono) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 888 | then have \<open>(sum (f \<circ> g) \<longlongrightarrow> f x) (finite_subsets_at_top S)\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 889 | using tendsto_cong f_sum | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 890 | by (simp add: Lim_transform_eventually eventually_finite_subsets_at_top_weakI) | 
| 77359 | 891 | then show \<open>((f \<circ> g) has_sum (f x)) S\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 892 | using has_sum_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 893 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 894 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 895 | lemma summable_on_comm_additive_general: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 896 |   fixes f :: \<open>'b :: {comm_monoid_add,topological_space} \<Rightarrow> 'c :: {comm_monoid_add,topological_space}\<close>
 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 897 | assumes \<open>\<And>F. finite F \<Longrightarrow> F \<subseteq> S \<Longrightarrow> sum (f \<circ> g) F = f (sum g F)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 898 | \<comment> \<open>Not using \<^const>\<open>additive\<close> because it would add sort constraint \<^class>\<open>ab_group_add\<close>\<close> | 
| 77359 | 899 | assumes \<open>\<And>x. (g has_sum x) S \<Longrightarrow> f \<midarrow>x\<rightarrow> f x\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 900 |     \<comment> \<open>For \<^class>\<open>t2_space\<close>, this is equivalent to \<open>isCont f x\<close> by @{thm [source] isCont_def}.\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 901 | assumes \<open>g summable_on S\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 902 | shows \<open>(f \<circ> g) summable_on S\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 903 | by (meson assms summable_on_def has_sum_comm_additive_general has_sum_def infsum_tendsto) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 904 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 905 | lemma infsum_comm_additive_general: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 906 |   fixes f :: \<open>'b :: {comm_monoid_add,t2_space} \<Rightarrow> 'c :: {comm_monoid_add,t2_space}\<close>
 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 907 | assumes f_sum: \<open>\<And>F. finite F \<Longrightarrow> F \<subseteq> S \<Longrightarrow> sum (f \<circ> g) F = f (sum g F)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 908 | \<comment> \<open>Not using \<^const>\<open>additive\<close> because it would add sort constraint \<^class>\<open>ab_group_add\<close>\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 909 | assumes \<open>isCont f (infsum g S)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 910 | assumes \<open>g summable_on S\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 911 | shows \<open>infsum (f \<circ> g) S = f (infsum g S)\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 912 | using assms | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 913 | by (intro infsumI has_sum_comm_additive_general has_sum_infsum) (auto simp: isCont_def) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 914 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 915 | lemma has_sum_comm_additive: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 916 |   fixes f :: \<open>'b :: {ab_group_add,topological_space} \<Rightarrow> 'c :: {ab_group_add,topological_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 917 | assumes \<open>additive f\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 918 | assumes \<open>f \<midarrow>x\<rightarrow> f x\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 919 |     \<comment> \<open>For \<^class>\<open>t2_space\<close>, this is equivalent to \<open>isCont f x\<close> by @{thm [source] isCont_def}.\<close>
 | 
| 77359 | 920 | assumes infsum: \<open>(g has_sum x) S\<close> | 
| 921 | shows \<open>((f \<circ> g) has_sum (f x)) S\<close> | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 922 | using assms | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 923 | by (intro has_sum_comm_additive_general has_sum_infsum) (auto simp: isCont_def additive.sum) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 924 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 925 | lemma summable_on_comm_additive: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 926 |   fixes f :: \<open>'b :: {ab_group_add,t2_space} \<Rightarrow> 'c :: {ab_group_add,topological_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 927 | assumes \<open>additive f\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 928 | assumes \<open>isCont f (infsum g S)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 929 | assumes \<open>g summable_on S\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 930 | shows \<open>(f \<circ> g) summable_on S\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 931 | by (meson assms summable_on_def has_sum_comm_additive has_sum_infsum isContD) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 932 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 933 | lemma infsum_comm_additive: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 934 |   fixes f :: \<open>'b :: {ab_group_add,t2_space} \<Rightarrow> 'c :: {ab_group_add,t2_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 935 | assumes \<open>additive f\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 936 | assumes \<open>isCont f (infsum g S)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 937 | assumes \<open>g summable_on S\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 938 | shows \<open>infsum (f \<circ> g) S = f (infsum g S)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 939 | by (rule infsum_comm_additive_general; auto simp: assms additive.sum) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 940 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 941 | lemma nonneg_bdd_above_has_sum: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 942 |   fixes f :: \<open>'a \<Rightarrow> 'b :: {conditionally_complete_linorder, ordered_comm_monoid_add, linorder_topology}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 943 | assumes \<open>\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 944 |   assumes \<open>bdd_above (sum f ` {F. F\<subseteq>A \<and> finite F})\<close>
 | 
| 77359 | 945 |   shows \<open>(f has_sum (SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F)) A\<close>
 | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 946 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 947 |   have \<open>(sum f \<longlongrightarrow> (SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F)) (finite_subsets_at_top A)\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 948 | proof (rule order_tendstoI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 949 |     fix a assume \<open>a < (SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F)\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 950 | then obtain F where \<open>a < sum f F\<close> and \<open>finite F\<close> and \<open>F \<subseteq> A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 951 | by (metis (mono_tags, lifting) Collect_cong Collect_empty_eq assms(2) empty_subsetI finite.emptyI less_cSUP_iff mem_Collect_eq) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 952 | have "\<And>Y. \<lbrakk>finite Y; F \<subseteq> Y; Y \<subseteq> A\<rbrakk> \<Longrightarrow> a < sum f Y" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 953 | by (meson DiffE \<open>a < sum f F\<close> assms(1) less_le_trans subset_iff sum_mono2) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 954 | then show \<open>\<forall>\<^sub>F x in finite_subsets_at_top A. a < sum f x\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 955 | by (metis \<open>F \<subseteq> A\<close> \<open>finite F\<close> eventually_finite_subsets_at_top) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 956 | next | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 957 |     fix a assume *: \<open>(SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F) < a\<close>
 | 
| 77388 | 958 |     have "sum f F \<le> (SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F)" if \<open>F\<subseteq>A\<close> and \<open>finite F\<close> for F
 | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 959 | by (rule cSUP_upper) (use that assms(2) in \<open>auto simp: conj_commute\<close>) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 960 | then show \<open>\<forall>\<^sub>F x in finite_subsets_at_top A. sum f x < a\<close> | 
| 77388 | 961 | by (metis (no_types, lifting) "*" eventually_finite_subsets_at_top_weakI order_le_less_trans) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 962 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 963 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 964 | using has_sum_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 965 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 966 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 967 | lemma nonneg_bdd_above_summable_on: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 968 |   fixes f :: \<open>'a \<Rightarrow> 'b :: {conditionally_complete_linorder, ordered_comm_monoid_add, linorder_topology}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 969 | assumes \<open>\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 970 |   assumes \<open>bdd_above (sum f ` {F. F\<subseteq>A \<and> finite F})\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 971 | shows \<open>f summable_on A\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 972 | using assms summable_on_def nonneg_bdd_above_has_sum by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 973 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 974 | lemma nonneg_bdd_above_infsum: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 975 |   fixes f :: \<open>'a \<Rightarrow> 'b :: {conditionally_complete_linorder, ordered_comm_monoid_add, linorder_topology}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 976 | assumes \<open>\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 977 |   assumes \<open>bdd_above (sum f ` {F. F\<subseteq>A \<and> finite F})\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 978 |   shows \<open>infsum f A = (SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F)\<close>
 | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 979 | using assms by (auto intro!: infsumI nonneg_bdd_above_has_sum) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 980 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 981 | lemma nonneg_has_sum_complete: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 982 |   fixes f :: \<open>'a \<Rightarrow> 'b :: {complete_linorder, ordered_comm_monoid_add, linorder_topology}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 983 | assumes \<open>\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0\<close> | 
| 77359 | 984 |   shows \<open>(f has_sum (SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F)) A\<close>
 | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 985 | using assms nonneg_bdd_above_has_sum by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 986 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 987 | lemma nonneg_summable_on_complete: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 988 |   fixes f :: \<open>'a \<Rightarrow> 'b :: {complete_linorder, ordered_comm_monoid_add, linorder_topology}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 989 | assumes \<open>\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 990 | shows \<open>f summable_on A\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 991 | using assms nonneg_bdd_above_summable_on by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 992 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 993 | lemma nonneg_infsum_complete: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 994 |   fixes f :: \<open>'a \<Rightarrow> 'b :: {complete_linorder, ordered_comm_monoid_add, linorder_topology}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 995 | assumes \<open>\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 996 |   shows \<open>infsum f A = (SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F)\<close>
 | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 997 | using assms nonneg_bdd_above_infsum by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 998 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 999 | lemma has_sum_nonneg: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1000 |   fixes f :: "'a \<Rightarrow> 'b::{ordered_comm_monoid_add,linorder_topology}"
 | 
| 77359 | 1001 | assumes "(f has_sum a) M" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1002 | and "\<And>x. x \<in> M \<Longrightarrow> 0 \<le> f x" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1003 | shows "a \<ge> 0" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1004 | by (metis (no_types, lifting) DiffD1 assms empty_iff has_sum_0 has_sum_mono_neutral order_refl) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1005 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1006 | lemma infsum_nonneg: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1007 |   fixes f :: "'a \<Rightarrow> 'b::{ordered_comm_monoid_add,linorder_topology}"
 | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1008 | assumes "\<And>x. x \<in> M \<Longrightarrow> 0 \<le> f x" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1009 | shows "infsum f M \<ge> 0" (is "?lhs \<ge> _") | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1010 | by (metis assms has_sum_infsum has_sum_nonneg infsum_not_exists linorder_linear) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1011 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1012 | lemma has_sum_mono2: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1013 |   fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}"
 | 
| 77359 | 1014 | assumes "(f has_sum S) A" "(f has_sum S') B" "A \<subseteq> B" | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1015 | assumes "\<And>x. x \<in> B - A \<Longrightarrow> f x \<ge> 0" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1016 | shows "S \<le> S'" | 
| 77388 | 1017 | by (metis add_0 add_right_mono assms diff_add_cancel has_sum_Diff has_sum_nonneg) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1018 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1019 | lemma infsum_mono2: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1020 |   fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}"
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1021 | assumes "f summable_on A" "f summable_on B" "A \<subseteq> B" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1022 | assumes "\<And>x. x \<in> B - A \<Longrightarrow> f x \<ge> 0" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1023 | shows "infsum f A \<le> infsum f B" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1024 | by (rule has_sum_mono2[OF has_sum_infsum has_sum_infsum]) (use assms in auto) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1025 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1026 | lemma finite_sum_le_has_sum: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1027 |   fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}"
 | 
| 77359 | 1028 | assumes "(f has_sum S) A" "finite B" "B \<subseteq> A" | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1029 | assumes "\<And>x. x \<in> A - B \<Longrightarrow> f x \<ge> 0" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1030 | shows "sum f B \<le> S" | 
| 77388 | 1031 | by (meson assms has_sum_finite has_sum_mono2) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1032 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1033 | lemma finite_sum_le_infsum: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1034 |   fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}"
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1035 | assumes "f summable_on A" "finite B" "B \<subseteq> A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1036 | assumes "\<And>x. x \<in> A - B \<Longrightarrow> f x \<ge> 0" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1037 | shows "sum f B \<le> infsum f A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1038 | by (rule finite_sum_le_has_sum[OF has_sum_infsum]) (use assms in auto) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1039 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1040 | lemma has_sum_reindex: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1041 | assumes \<open>inj_on h A\<close> | 
| 77359 | 1042 | shows \<open>(g has_sum x) (h ` A) \<longleftrightarrow> ((g \<circ> h) has_sum x) A\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1043 | proof - | 
| 77359 | 1044 | have \<open>(g has_sum x) (h ` A) \<longleftrightarrow> (sum g \<longlongrightarrow> x) (finite_subsets_at_top (h ` A))\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1045 | by (simp add: has_sum_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1046 | also have \<open>\<dots> \<longleftrightarrow> ((\<lambda>F. sum g (h ` F)) \<longlongrightarrow> x) (finite_subsets_at_top A)\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1047 | by (metis assms filterlim_filtermap filtermap_image_finite_subsets_at_top) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1048 | also have \<open>\<dots> \<longleftrightarrow> (sum (g \<circ> h) \<longlongrightarrow> x) (finite_subsets_at_top A)\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1049 | proof (intro tendsto_cong eventually_finite_subsets_at_top_weakI sum.reindex) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1050 | show "\<And>X. \<lbrakk>finite X; X \<subseteq> A\<rbrakk> \<Longrightarrow> inj_on h X" | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1051 | using assms subset_inj_on by blast | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1052 | qed | 
| 77359 | 1053 | also have \<open>\<dots> \<longleftrightarrow> ((g \<circ> h) has_sum x) A\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1054 | by (simp add: has_sum_def) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1055 | finally show ?thesis . | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1056 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1057 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1058 | lemma summable_on_reindex: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1059 | assumes \<open>inj_on h A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1060 | shows \<open>g summable_on (h ` A) \<longleftrightarrow> (g \<circ> h) summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1061 | by (simp add: assms summable_on_def has_sum_reindex) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1062 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1063 | lemma infsum_reindex: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1064 | assumes \<open>inj_on h A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1065 | shows \<open>infsum g (h ` A) = infsum (g \<circ> h) A\<close> | 
| 77388 | 1066 | by (metis assms has_sum_infsum has_sum_reindex infsumI infsum_def) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1067 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1068 | lemma summable_on_reindex_bij_betw: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1069 | assumes "bij_betw g A B" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1070 | shows "(\<lambda>x. f (g x)) summable_on A \<longleftrightarrow> f summable_on B" | 
| 77388 | 1071 | by (smt (verit) assms bij_betw_def o_apply summable_on_cong summable_on_reindex) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1072 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1073 | lemma infsum_reindex_bij_betw: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1074 | assumes "bij_betw g A B" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1075 | shows "infsum (\<lambda>x. f (g x)) A = infsum f B" | 
| 77388 | 1076 | by (metis (mono_tags, lifting) assms bij_betw_def infsum_cong infsum_reindex o_def) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1077 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1078 | lemma sum_uniformity: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1079 |   assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'b::{uniform_space,comm_monoid_add},y). x+y)\<close>
 | 
| 77388 | 1080 | assumes EE: \<open>eventually E uniformity\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1081 | obtains D where \<open>eventually D uniformity\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1082 | and \<open>\<And>M::'a set. \<And>f f' :: 'a \<Rightarrow> 'b. card M \<le> n \<and> (\<forall>m\<in>M. D (f m, f' m)) \<Longrightarrow> E (sum f M, sum f' M)\<close> | 
| 77388 | 1083 | proof (atomize_elim, insert EE, induction n arbitrary: E rule:nat_induct) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1084 | case 0 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1085 | then show ?case | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1086 | by (metis card_eq_0_iff equals0D le_zero_eq sum.infinite sum.not_neutral_contains_not_neutral uniformity_refl) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1087 | next | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1088 | case (Suc n) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1089 | from plus_cont[unfolded uniformly_continuous_on_uniformity filterlim_def le_filter_def, rule_format, OF Suc.prems] | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1090 | obtain D1 D2 where \<open>eventually D1 uniformity\<close> and \<open>eventually D2 uniformity\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1091 | and D1D2E: \<open>D1 (x, y) \<Longrightarrow> D2 (x', y') \<Longrightarrow> E (x + x', y + y')\<close> for x y x' y' | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1092 | apply atomize_elim | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1093 | by (auto simp: eventually_prod_filter case_prod_beta uniformity_prod_def eventually_filtermap) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1094 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1095 | from Suc.IH[OF \<open>eventually D2 uniformity\<close>] | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1096 | obtain D3 where \<open>eventually D3 uniformity\<close> and D3: \<open>card M \<le> n \<Longrightarrow> (\<forall>m\<in>M. D3 (f m, f' m)) \<Longrightarrow> D2 (sum f M, sum f' M)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1097 | for M :: \<open>'a set\<close> and f f' | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1098 | by metis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1099 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1100 | define D where \<open>D x \<equiv> D1 x \<and> D3 x\<close> for x | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1101 | have \<open>eventually D uniformity\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1102 | using D_def \<open>eventually D1 uniformity\<close> \<open>eventually D3 uniformity\<close> eventually_elim2 by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1103 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1104 | have \<open>E (sum f M, sum f' M)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1105 | if \<open>card M \<le> Suc n\<close> and DM: \<open>\<forall>m\<in>M. D (f m, f' m)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1106 | for M :: \<open>'a set\<close> and f f' | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1107 | proof (cases \<open>card M = 0\<close>) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1108 | case True | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1109 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1110 | by (metis Suc.prems card_eq_0_iff sum.empty sum.infinite uniformity_refl) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1111 | next | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1112 | case False | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1113 | with \<open>card M \<le> Suc n\<close> obtain N x where \<open>card N \<le> n\<close> and \<open>x \<notin> N\<close> and \<open>M = insert x N\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1114 | by (metis card_Suc_eq less_Suc_eq_0_disj less_Suc_eq_le) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1115 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1116 | from DM have \<open>\<And>m. m\<in>N \<Longrightarrow> D (f m, f' m)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1117 | using \<open>M = insert x N\<close> by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1118 | with D3[OF \<open>card N \<le> n\<close>] | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1119 | have D2_N: \<open>D2 (sum f N, sum f' N)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1120 | using D_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1121 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1122 | from DM | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1123 | have \<open>D (f x, f' x)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1124 | using \<open>M = insert x N\<close> by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1125 | then have \<open>D1 (f x, f' x)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1126 | by (simp add: D_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1127 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1128 | with D2_N | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1129 | have \<open>E (f x + sum f N, f' x + sum f' N)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1130 | using D1D2E by presburger | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1131 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1132 | then show \<open>E (sum f M, sum f' M)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1133 | by (metis False \<open>M = insert x N\<close> \<open>x \<notin> N\<close> card.infinite finite_insert sum.insert) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1134 | qed | 
| 77388 | 1135 | with \<open>eventually D uniformity\<close> show ?case | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1136 | by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1137 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1138 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1139 | lemma has_sum_Sigma: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1140 | fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1141 |     and f :: \<open>'a \<times> 'b \<Rightarrow> 'c::{comm_monoid_add,uniform_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1142 | assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'c,y). x+y)\<close> | 
| 77359 | 1143 | assumes summableAB: "(f has_sum a) (Sigma A B)" | 
| 1144 | assumes summableB: \<open>\<And>x. x\<in>A \<Longrightarrow> ((\<lambda>y. f (x, y)) has_sum b x) (B x)\<close> | |
| 1145 | shows "(b has_sum a) A" | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1146 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1147 | define F FB FA where \<open>F = finite_subsets_at_top (Sigma A B)\<close> and \<open>FB x = finite_subsets_at_top (B x)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1148 | and \<open>FA = finite_subsets_at_top A\<close> for x | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1149 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1150 | from summableB | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1151 | have sum_b: \<open>(sum (\<lambda>y. f (x, y)) \<longlongrightarrow> b x) (FB x)\<close> if \<open>x \<in> A\<close> for x | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1152 | using FB_def[abs_def] has_sum_def that by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1153 | from summableAB | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1154 | have sum_S: \<open>(sum f \<longlongrightarrow> a) F\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1155 | using F_def has_sum_def by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1156 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1157 |   have finite_proj: \<open>finite {b| b. (a,b) \<in> H}\<close> if \<open>finite H\<close> for H :: \<open>('a\<times>'b) set\<close> and a
 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1158 | by (metis (no_types, lifting) finite_imageI finite_subset image_eqI mem_Collect_eq snd_conv subsetI that) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1159 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1160 | have \<open>(sum b \<longlongrightarrow> a) FA\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1161 | proof (rule tendsto_iff_uniformity[THEN iffD2, rule_format]) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1162 |     fix E :: \<open>('c \<times> 'c) \<Rightarrow> bool\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1163 | assume \<open>eventually E uniformity\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1164 | then obtain D where D_uni: \<open>eventually D uniformity\<close> and DDE': \<open>\<And>x y z. D (x, y) \<Longrightarrow> D (y, z) \<Longrightarrow> E (x, z)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1165 | by (metis (no_types, lifting) \<open>eventually E uniformity\<close> uniformity_transE) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1166 | from sum_S obtain G where \<open>finite G\<close> and \<open>G \<subseteq> Sigma A B\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1167 | and G_sum: \<open>G \<subseteq> H \<Longrightarrow> H \<subseteq> Sigma A B \<Longrightarrow> finite H \<Longrightarrow> D (sum f H, a)\<close> for H | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1168 | unfolding tendsto_iff_uniformity | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1169 | by (metis (mono_tags, lifting) D_uni F_def eventually_finite_subsets_at_top) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1170 | have \<open>finite (fst ` G)\<close> and \<open>fst ` G \<subseteq> A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1171 | using \<open>finite G\<close> \<open>G \<subseteq> Sigma A B\<close> by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1172 | thm uniformity_prod_def | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1173 |     define Ga where \<open>Ga a = {b. (a,b) \<in> G}\<close> for a
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1174 | have Ga_fin: \<open>finite (Ga a)\<close> and Ga_B: \<open>Ga a \<subseteq> B a\<close> for a | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1175 | using \<open>finite G\<close> \<open>G \<subseteq> Sigma A B\<close> finite_proj by (auto simp: Ga_def finite_proj) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1176 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1177 | have \<open>E (sum b M, a)\<close> if \<open>M \<supseteq> fst ` G\<close> and \<open>finite M\<close> and \<open>M \<subseteq> A\<close> for M | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1178 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1179 | define FMB where \<open>FMB = finite_subsets_at_top (Sigma M B)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1180 | have \<open>eventually (\<lambda>H. D (\<Sum>a\<in>M. b a, \<Sum>(a,b)\<in>H. f (a,b))) FMB\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1181 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1182 | obtain D' where D'_uni: \<open>eventually D' uniformity\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1183 | and \<open>card M' \<le> card M \<and> (\<forall>m\<in>M'. D' (g m, g' m)) \<Longrightarrow> D (sum g M', sum g' M')\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1184 | for M' :: \<open>'a set\<close> and g g' | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1185 | using sum_uniformity[OF plus_cont \<open>eventually D uniformity\<close>] by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1186 | then have D'_sum_D: \<open>(\<forall>m\<in>M. D' (g m, g' m)) \<Longrightarrow> D (sum g M, sum g' M)\<close> for g g' | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1187 | by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1188 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1189 | obtain Ha where \<open>Ha a \<supseteq> Ga a\<close> and Ha_fin: \<open>finite (Ha a)\<close> and Ha_B: \<open>Ha a \<subseteq> B a\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1190 | and D'_sum_Ha: \<open>Ha a \<subseteq> L \<Longrightarrow> L \<subseteq> B a \<Longrightarrow> finite L \<Longrightarrow> D' (b a, sum (\<lambda>b. f (a,b)) L)\<close> if \<open>a \<in> A\<close> for a L | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1191 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1192 | from sum_b[unfolded tendsto_iff_uniformity, rule_format, OF _ D'_uni[THEN uniformity_sym]] | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1193 | obtain Ha0 where \<open>finite (Ha0 a)\<close> and \<open>Ha0 a \<subseteq> B a\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1194 | and \<open>Ha0 a \<subseteq> L \<Longrightarrow> L \<subseteq> B a \<Longrightarrow> finite L \<Longrightarrow> D' (b a, sum (\<lambda>b. f (a,b)) L)\<close> if \<open>a \<in> A\<close> for a L | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1195 | unfolding FB_def eventually_finite_subsets_at_top unfolding prod.case by metis | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1196 | moreover define Ha where \<open>Ha a = Ha0 a \<union> Ga a\<close> for a | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1197 | ultimately show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1198 | using that[where Ha=Ha] | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1199 | using Ga_fin Ga_B by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1200 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1201 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1202 | have \<open>D (\<Sum>a\<in>M. b a, \<Sum>(a,b)\<in>H. f (a,b))\<close> if \<open>finite H\<close> and \<open>H \<subseteq> Sigma M B\<close> and \<open>H \<supseteq> Sigma M Ha\<close> for H | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1203 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1204 |           define Ha' where \<open>Ha' a = {b| b. (a,b) \<in> H}\<close> for a
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1205 | have [simp]: \<open>finite (Ha' a)\<close> and [simp]: \<open>Ha' a \<supseteq> Ha a\<close> and [simp]: \<open>Ha' a \<subseteq> B a\<close> if \<open>a \<in> M\<close> for a | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1206 | unfolding Ha'_def using \<open>finite H\<close> \<open>H \<subseteq> Sigma M B\<close> \<open>Sigma M Ha \<subseteq> H\<close> that finite_proj by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1207 | have \<open>Sigma M Ha' = H\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1208 | using that by (auto simp: Ha'_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1209 | then have *: \<open>(\<Sum>(a,b)\<in>H. f (a,b)) = (\<Sum>a\<in>M. \<Sum>b\<in>Ha' a. f (a,b))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1210 | by (simp add: \<open>finite M\<close> sum.Sigma) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1211 | have \<open>D' (b a, sum (\<lambda>b. f (a,b)) (Ha' a))\<close> if \<open>a \<in> M\<close> for a | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1212 | using D'_sum_Ha \<open>M \<subseteq> A\<close> that by auto | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1213 | then have \<open>D (\<Sum>a\<in>M. b a, \<Sum>a\<in>M. sum (\<lambda>b. f (a,b)) (Ha' a))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1214 | by (rule_tac D'_sum_D, auto) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1215 | with * show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1216 | by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1217 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1218 | moreover have \<open>Sigma M Ha \<subseteq> Sigma M B\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1219 | using Ha_B \<open>M \<subseteq> A\<close> by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1220 | ultimately show ?thesis | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1221 | unfolding FMB_def eventually_finite_subsets_at_top | 
| 77388 | 1222 | by (metis (no_types, lifting) Ha_fin finite_SigmaI subsetD that(2) that(3)) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1223 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1224 | moreover have \<open>eventually (\<lambda>H. D (\<Sum>(a,b)\<in>H. f (a,b), a)) FMB\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1225 | unfolding FMB_def eventually_finite_subsets_at_top | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1226 | proof (rule exI[of _ G], safe) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1227 | fix Y assume Y: "finite Y" "G \<subseteq> Y" "Y \<subseteq> Sigma M B" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1228 | thus "D (\<Sum>(a,b)\<in>Y. f (a, b), a)" | 
| 77388 | 1229 | using G_sum[of Y] Y using that(3) by fastforce | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1230 | qed (use \<open>finite G\<close> \<open>G \<subseteq> Sigma A B\<close> that in auto) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1231 | ultimately have \<open>\<forall>\<^sub>F x in FMB. E (sum b M, a)\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1232 | by eventually_elim (use DDE' in auto) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1233 | then show \<open>E (sum b M, a)\<close> | 
| 77388 | 1234 | using FMB_def by force | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1235 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1236 | then show \<open>\<forall>\<^sub>F x in FA. E (sum b x, a)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1237 | using \<open>finite (fst ` G)\<close> and \<open>fst ` G \<subseteq> A\<close> | 
| 77388 | 1238 | by (metis (mono_tags, lifting) FA_def eventually_finite_subsets_at_top) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1239 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1240 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1241 | by (simp add: FA_def has_sum_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1242 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1243 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1244 | lemma summable_on_Sigma: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1245 | fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1246 |     and f :: \<open>'a \<Rightarrow> 'b \<Rightarrow> 'c::{comm_monoid_add, t2_space, uniform_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1247 | assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'c,y). x+y)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1248 | assumes summableAB: "(\<lambda>(x,y). f x y) summable_on (Sigma A B)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1249 | assumes summableB: \<open>\<And>x. x\<in>A \<Longrightarrow> (f x) summable_on (B x)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1250 | shows \<open>(\<lambda>x. infsum (f x) (B x)) summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1251 | proof - | 
| 77359 | 1252 | from summableAB obtain a where a: \<open>((\<lambda>(x,y). f x y) has_sum a) (Sigma A B)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1253 | using has_sum_infsum by blast | 
| 77359 | 1254 | from summableB have b: \<open>\<And>x. x\<in>A \<Longrightarrow> (f x has_sum infsum (f x) (B x)) (B x)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1255 | by (auto intro!: has_sum_infsum) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1256 | show ?thesis | 
| 77388 | 1257 | using plus_cont a b | 
| 1258 | by (smt (verit) has_sum_Sigma[where f=\<open>\<lambda>(x,y). f x y\<close>] has_sum_cong old.prod.case summable_on_def) | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1259 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1260 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1261 | lemma infsum_Sigma: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1262 | fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1263 |     and f :: \<open>'a \<times> 'b \<Rightarrow> 'c::{comm_monoid_add, t2_space, uniform_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1264 | assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'c,y). x+y)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1265 | assumes summableAB: "f summable_on (Sigma A B)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1266 | assumes summableB: \<open>\<And>x. x\<in>A \<Longrightarrow> (\<lambda>y. f (x, y)) summable_on (B x)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1267 | shows "infsum f (Sigma A B) = infsum (\<lambda>x. infsum (\<lambda>y. f (x, y)) (B x)) A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1268 | proof - | 
| 77359 | 1269 | from summableAB have a: \<open>(f has_sum infsum f (Sigma A B)) (Sigma A B)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1270 | using has_sum_infsum by blast | 
| 77359 | 1271 | from summableB have b: \<open>\<And>x. x\<in>A \<Longrightarrow> ((\<lambda>y. f (x, y)) has_sum infsum (\<lambda>y. f (x, y)) (B x)) (B x)\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1272 | by (auto intro!: has_sum_infsum) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1273 | show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1274 | using plus_cont a b by (auto intro: infsumI[symmetric] has_sum_Sigma simp: summable_on_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1275 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1276 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1277 | lemma infsum_Sigma': | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1278 | fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1279 |     and f :: \<open>'a \<Rightarrow> 'b \<Rightarrow> 'c::{comm_monoid_add, t2_space, uniform_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1280 | assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'c,y). x+y)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1281 | assumes summableAB: "(\<lambda>(x,y). f x y) summable_on (Sigma A B)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1282 | assumes summableB: \<open>\<And>x. x\<in>A \<Longrightarrow> (f x) summable_on (B x)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1283 | shows \<open>infsum (\<lambda>x. infsum (f x) (B x)) A = infsum (\<lambda>(x,y). f x y) (Sigma A B)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1284 | using infsum_Sigma[of \<open>\<lambda>(x,y). f x y\<close> A B] | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1285 | using assms by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1286 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1287 | text \<open>A special case of @{thm [source] infsum_Sigma} etc. for Banach spaces. It has less premises.\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1288 | lemma | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1289 | fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1290 | and f :: \<open>'a \<Rightarrow> 'b \<Rightarrow> 'c::banach\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1291 | assumes [simp]: "(\<lambda>(x,y). f x y) summable_on (Sigma A B)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1292 | shows infsum_Sigma'_banach: \<open>infsum (\<lambda>x. infsum (f x) (B x)) A = infsum (\<lambda>(x,y). f x y) (Sigma A B)\<close> (is ?thesis1) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1293 | and summable_on_Sigma_banach: \<open>(\<lambda>x. infsum (f x) (B x)) summable_on A\<close> (is ?thesis2) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1294 | proof - | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1295 | have fsum: \<open>(f x) summable_on (B x)\<close> if \<open>x \<in> A\<close> for x | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1296 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1297 | from assms | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1298 | have \<open>(\<lambda>(x,y). f x y) summable_on (Pair x ` B x)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1299 | by (meson image_subset_iff summable_on_subset_banach mem_Sigma_iff that) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1300 | then have \<open>((\<lambda>(x,y). f x y) \<circ> Pair x) summable_on (B x)\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1301 | by (metis summable_on_reindex inj_on_def prod.inject) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1302 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1303 | by (auto simp: o_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1304 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1305 | show ?thesis1 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1306 | using fsum assms infsum_Sigma' isUCont_plus by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1307 | show ?thesis2 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1308 | using fsum assms isUCont_plus summable_on_Sigma by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1309 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1310 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1311 | lemma infsum_Sigma_banach: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1312 | fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1313 | and f :: \<open>'a \<times> 'b \<Rightarrow> 'c::banach\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1314 | assumes [simp]: "f summable_on (Sigma A B)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1315 | shows \<open>infsum (\<lambda>x. infsum (\<lambda>y. f (x,y)) (B x)) A = infsum f (Sigma A B)\<close> | 
| 77388 | 1316 | using assms by (simp add: infsum_Sigma'_banach) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1317 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1318 | lemma infsum_swap: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1319 | fixes A :: "'a set" and B :: "'b set" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1320 |   fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::{comm_monoid_add,t2_space,uniform_space}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1321 | assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'c,y). x+y)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1322 | assumes \<open>(\<lambda>(x, y). f x y) summable_on (A \<times> B)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1323 | assumes \<open>\<And>a. a\<in>A \<Longrightarrow> (f a) summable_on B\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1324 | assumes \<open>\<And>b. b\<in>B \<Longrightarrow> (\<lambda>a. f a b) summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1325 | shows \<open>infsum (\<lambda>x. infsum (\<lambda>y. f x y) B) A = infsum (\<lambda>y. infsum (\<lambda>x. f x y) A) B\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1326 | proof - | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1327 | have "(\<lambda>(x, y). f y x) \<circ> prod.swap summable_on A \<times> B" | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1328 | by (simp add: assms(2) summable_on_cong) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1329 | then have fyx: \<open>(\<lambda>(x, y). f y x) summable_on (B \<times> A)\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1330 | by (metis has_sum_reindex infsum_reindex inj_swap product_swap summable_iff_has_sum_infsum) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1331 | have \<open>infsum (\<lambda>x. infsum (\<lambda>y. f x y) B) A = infsum (\<lambda>(x,y). f x y) (A \<times> B)\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1332 | using assms infsum_Sigma' by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1333 | also have \<open>\<dots> = infsum (\<lambda>(x,y). f y x) (B \<times> A)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1334 | apply (subst product_swap[symmetric]) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1335 | apply (subst infsum_reindex) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1336 | using assms by (auto simp: o_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1337 | also have \<open>\<dots> = infsum (\<lambda>y. infsum (\<lambda>x. f x y) A) B\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1338 | by (smt (verit) fyx assms(1) assms(4) infsum_Sigma' infsum_cong) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1339 | finally show ?thesis . | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1340 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1341 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1342 | lemma infsum_swap_banach: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1343 | fixes A :: "'a set" and B :: "'b set" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1344 | fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::banach" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1345 | assumes \<open>(\<lambda>(x, y). f x y) summable_on (A \<times> B)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1346 | shows "infsum (\<lambda>x. infsum (\<lambda>y. f x y) B) A = infsum (\<lambda>y. infsum (\<lambda>x. f x y) A) B" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1347 | proof - | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1348 | have \<section>: \<open>(\<lambda>(x, y). f y x) summable_on (B \<times> A)\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1349 | by (metis (mono_tags, lifting) assms case_swap inj_swap o_apply product_swap summable_on_cong summable_on_reindex) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1350 | have \<open>infsum (\<lambda>x. infsum (\<lambda>y. f x y) B) A = infsum (\<lambda>(x,y). f x y) (A \<times> B)\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1351 | using assms infsum_Sigma'_banach by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1352 | also have \<open>\<dots> = infsum (\<lambda>(x,y). f y x) (B \<times> A)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1353 | apply (subst product_swap[symmetric]) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1354 | apply (subst infsum_reindex) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1355 | using assms by (auto simp: o_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1356 | also have \<open>\<dots> = infsum (\<lambda>y. infsum (\<lambda>x. f x y) A) B\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1357 | by (metis (mono_tags, lifting) \<section> infsum_Sigma'_banach infsum_cong) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1358 | finally show ?thesis . | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1359 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1360 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1361 | lemma nonneg_infsum_le_0D: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1362 |   fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add,ordered_ab_group_add,linorder_topology}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1363 | assumes "infsum f A \<le> 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1364 | and abs_sum: "f summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1365 | and nneg: "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1366 | and "x \<in> A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1367 | shows "f x = 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1368 | proof (rule ccontr) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1369 | assume \<open>f x \<noteq> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1370 |   have ex: \<open>f summable_on (A-{x})\<close>
 | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1371 | by (rule summable_on_cofin_subset) (use assms in auto) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1372 |   have pos: \<open>infsum f (A - {x}) \<ge> 0\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1373 | by (rule infsum_nonneg) (use nneg in auto) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1374 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1375 | have [trans]: \<open>x \<ge> y \<Longrightarrow> y > z \<Longrightarrow> x > z\<close> for x y z :: 'b by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1376 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1377 |   have \<open>infsum f A = infsum f (A-{x}) + infsum f {x}\<close>
 | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1378 | by (subst infsum_Un_disjoint[symmetric]) (use assms ex in \<open>auto simp: insert_absorb\<close>) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1379 |   also have \<open>\<dots> \<ge> infsum f {x}\<close> (is \<open>_ \<ge> \<dots>\<close>)
 | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1380 | using pos by (rule add_increasing) simp | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1381 | also have \<open>\<dots> = f x\<close> (is \<open>_ = \<dots>\<close>) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1382 | by (subst infsum_finite) auto | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1383 | also have \<open>\<dots> > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1384 | using \<open>f x \<noteq> 0\<close> assms(4) nneg by fastforce | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1385 | finally show False | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1386 | using assms by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1387 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1388 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1389 | lemma nonneg_has_sum_le_0D: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1390 |   fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add,ordered_ab_group_add,linorder_topology}"
 | 
| 77359 | 1391 | assumes "(f has_sum a) A" \<open>a \<le> 0\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1392 | and "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> 0" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1393 | and "x \<in> A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1394 | shows "f x = 0" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1395 | by (metis assms infsumI nonneg_infsum_le_0D summable_on_def) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1396 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1397 | lemma has_sum_cmult_left: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1398 |   fixes f :: "'a \<Rightarrow> 'b :: {topological_semigroup_mult, semiring_0}"
 | 
| 77359 | 1399 | assumes \<open>(f has_sum a) A\<close> | 
| 1400 | shows "((\<lambda>x. f x * c) has_sum (a * c)) A" | |
| 77388 | 1401 | using assms tendsto_mult_right | 
| 1402 | by (force simp add: has_sum_def sum_distrib_right) | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1403 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1404 | lemma infsum_cmult_left: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1405 |   fixes f :: "'a \<Rightarrow> 'b :: {t2_space, topological_semigroup_mult, semiring_0}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1406 | assumes \<open>c \<noteq> 0 \<Longrightarrow> f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1407 | shows "infsum (\<lambda>x. f x * c) A = infsum f A * c" | 
| 77388 | 1408 | using assms has_sum_cmult_left infsumI summable_iff_has_sum_infsum by fastforce | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1409 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1410 | lemma summable_on_cmult_left: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1411 |   fixes f :: "'a \<Rightarrow> 'b :: {t2_space, topological_semigroup_mult, semiring_0}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1412 | assumes \<open>f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1413 | shows "(\<lambda>x. f x * c) summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1414 | using assms summable_on_def has_sum_cmult_left by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1415 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1416 | lemma has_sum_cmult_right: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1417 |   fixes f :: "'a \<Rightarrow> 'b :: {topological_semigroup_mult, semiring_0}"
 | 
| 77359 | 1418 | assumes \<open>(f has_sum a) A\<close> | 
| 1419 | shows "((\<lambda>x. c * f x) has_sum (c * a)) A" | |
| 77388 | 1420 | using assms tendsto_mult_left | 
| 1421 | by (force simp add: has_sum_def sum_distrib_left) | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1422 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1423 | lemma infsum_cmult_right: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1424 |   fixes f :: "'a \<Rightarrow> 'b :: {t2_space, topological_semigroup_mult, semiring_0}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1425 | assumes \<open>c \<noteq> 0 \<Longrightarrow> f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1426 | shows \<open>infsum (\<lambda>x. c * f x) A = c * infsum f A\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1427 | using assms has_sum_cmult_right infsumI summable_iff_has_sum_infsum by fastforce | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1428 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1429 | lemma summable_on_cmult_right: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1430 |   fixes f :: "'a \<Rightarrow> 'b :: {t2_space, topological_semigroup_mult, semiring_0}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1431 | assumes \<open>f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1432 | shows "(\<lambda>x. c * f x) summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1433 | using assms summable_on_def has_sum_cmult_right by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1434 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1435 | lemma summable_on_cmult_left': | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1436 |   fixes f :: "'a \<Rightarrow> 'b :: {t2_space, topological_semigroup_mult, division_ring}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1437 | assumes \<open>c \<noteq> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1438 | shows "(\<lambda>x. f x * c) summable_on A \<longleftrightarrow> f summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1439 | proof | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1440 | assume \<open>f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1441 | then show \<open>(\<lambda>x. f x * c) summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1442 | by (rule summable_on_cmult_left) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1443 | next | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1444 | assume \<open>(\<lambda>x. f x * c) summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1445 | then have \<open>(\<lambda>x. f x * c * inverse c) summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1446 | by (rule summable_on_cmult_left) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1447 | then show \<open>f summable_on A\<close> | 
| 77388 | 1448 | by (smt (verit, del_insts) assms divide_inverse nonzero_divide_eq_eq summable_on_cong) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1449 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1450 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1451 | lemma summable_on_cmult_right': | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1452 |   fixes f :: "'a \<Rightarrow> 'b :: {t2_space, topological_semigroup_mult, division_ring}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1453 | assumes \<open>c \<noteq> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1454 | shows "(\<lambda>x. c * f x) summable_on A \<longleftrightarrow> f summable_on A" | 
| 77388 | 1455 | by (metis (no_types, lifting) assms left_inverse mult.assoc mult_1 summable_on_cmult_right summable_on_cong) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1456 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1457 | lemma infsum_cmult_left': | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1458 |   fixes f :: "'a \<Rightarrow> 'b :: {t2_space, topological_semigroup_mult, division_ring}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1459 | shows "infsum (\<lambda>x. f x * c) A = infsum f A * c" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1460 | by (metis (full_types) infsum_cmult_left infsum_not_exists mult_eq_0_iff summable_on_cmult_left') | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1461 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1462 | lemma infsum_cmult_right': | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1463 |   fixes f :: "'a \<Rightarrow> 'b :: {t2_space,topological_semigroup_mult,division_ring}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1464 | shows "infsum (\<lambda>x. c * f x) A = c * infsum f A" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1465 | by (metis (full_types) infsum_cmult_right infsum_not_exists mult_eq_0_iff summable_on_cmult_right') | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1466 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1467 | lemma has_sum_constant[simp]: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1468 | assumes \<open>finite F\<close> | 
| 77359 | 1469 | shows \<open>((\<lambda>_. c) has_sum of_nat (card F) * c) F\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1470 | by (metis assms has_sum_finite sum_constant) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1471 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1472 | lemma infsum_constant[simp]: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1473 | assumes \<open>finite F\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1474 | shows \<open>infsum (\<lambda>_. c) F = of_nat (card F) * c\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1475 | by (simp add: assms) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1476 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1477 | lemma infsum_diverge_constant: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1478 | \<comment> \<open>This probably does not really need all of \<^class>\<open>archimedean_field\<close> but Isabelle/HOL | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1479 | has no type class such as, e.g., "archimedean ring".\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1480 |   fixes c :: \<open>'a::{archimedean_field, comm_monoid_add, linorder_topology, topological_semigroup_mult}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1481 | assumes \<open>infinite A\<close> and \<open>c \<noteq> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1482 | shows \<open>\<not> (\<lambda>_. c) summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1483 | proof (rule notI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1484 | assume \<open>(\<lambda>_. c) summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1485 | then have \<open>(\<lambda>_. inverse c * c) summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1486 | by (rule summable_on_cmult_right) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1487 | then have [simp]: \<open>(\<lambda>_. 1::'a) summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1488 | using assms by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1489 | have \<open>infsum (\<lambda>_. 1) A \<ge> d\<close> for d :: 'a | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1490 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1491 | obtain n :: nat where \<open>of_nat n \<ge> d\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1492 | by (meson real_arch_simple) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1493 | from assms | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1494 | obtain F where \<open>F \<subseteq> A\<close> and \<open>finite F\<close> and \<open>card F = n\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1495 | by (meson infinite_arbitrarily_large) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1496 | note \<open>d \<le> of_nat n\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1497 | also have \<open>of_nat n = infsum (\<lambda>_. 1::'a) F\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1498 | by (simp add: \<open>card F = n\<close> \<open>finite F\<close>) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1499 | also have \<open>\<dots> \<le> infsum (\<lambda>_. 1::'a) A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1500 | apply (rule infsum_mono_neutral) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1501 | using \<open>finite F\<close> \<open>F \<subseteq> A\<close> by auto | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1502 | finally show ?thesis . | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1503 | qed | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1504 | then show False | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1505 | by (meson linordered_field_no_ub not_less) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1506 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1507 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1508 | lemma has_sum_constant_archimedean[simp]: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1509 | \<comment> \<open>This probably does not really need all of \<^class>\<open>archimedean_field\<close> but Isabelle/HOL | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1510 | has no type class such as, e.g., "archimedean ring".\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1511 |   fixes c :: \<open>'a::{archimedean_field, comm_monoid_add, linorder_topology, topological_semigroup_mult}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1512 | shows \<open>infsum (\<lambda>_. c) A = of_nat (card A) * c\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1513 | by (metis infsum_0 infsum_constant infsum_diverge_constant infsum_not_exists sum.infinite sum_constant) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1514 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1515 | lemma has_sum_uminus: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1516 | fixes f :: \<open>'a \<Rightarrow> 'b::topological_ab_group_add\<close> | 
| 77359 | 1517 | shows \<open>((\<lambda>x. - f x) has_sum a) A \<longleftrightarrow> (f has_sum (- a)) A\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1518 | by (auto simp add: sum_negf[abs_def] tendsto_minus_cancel_left has_sum_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1519 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1520 | lemma summable_on_uminus: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1521 | fixes f :: \<open>'a \<Rightarrow> 'b::topological_ab_group_add\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1522 | shows\<open>(\<lambda>x. - f x) summable_on A \<longleftrightarrow> f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1523 | by (metis summable_on_def has_sum_uminus verit_minus_simplify(4)) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1524 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1525 | lemma infsum_uminus: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1526 |   fixes f :: \<open>'a \<Rightarrow> 'b::{topological_ab_group_add, t2_space}\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1527 | shows \<open>infsum (\<lambda>x. - f x) A = - infsum f A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1528 | by (metis (full_types) add.inverse_inverse add.inverse_neutral infsumI infsum_def has_sum_infsum has_sum_uminus) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1529 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1530 | lemma has_sum_le_finite_sums: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1531 |   fixes a :: \<open>'a::{comm_monoid_add,topological_space,linorder_topology}\<close>
 | 
| 77359 | 1532 | assumes \<open>(f has_sum a) A\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1533 | assumes \<open>\<And>F. finite F \<Longrightarrow> F \<subseteq> A \<Longrightarrow> sum f F \<le> b\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1534 | shows \<open>a \<le> b\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1535 | by (metis assms eventually_finite_subsets_at_top_weakI finite_subsets_at_top_neq_bot has_sum_def tendsto_upperbound) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1536 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1537 | lemma infsum_le_finite_sums: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1538 |   fixes b :: \<open>'a::{comm_monoid_add,topological_space,linorder_topology}\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1539 | assumes \<open>f summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1540 | assumes \<open>\<And>F. finite F \<Longrightarrow> F \<subseteq> A \<Longrightarrow> sum f F \<le> b\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1541 | shows \<open>infsum f A \<le> b\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1542 | by (meson assms has_sum_infsum has_sum_le_finite_sums) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1543 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1544 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1545 | lemma summable_on_scaleR_left [intro]: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1546 | fixes c :: \<open>'a :: real_normed_vector\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1547 | assumes "c \<noteq> 0 \<Longrightarrow> f summable_on A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1548 | shows "(\<lambda>x. f x *\<^sub>R c) summable_on A" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1549 | proof (cases \<open>c = 0\<close>) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1550 | case False | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1551 | then have "(\<lambda>y. y *\<^sub>R c) \<circ> f summable_on A" | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1552 | using assms by (auto simp add: scaleR_left.additive_axioms summable_on_comm_additive) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1553 | then show ?thesis | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1554 | by (metis (mono_tags, lifting) comp_apply summable_on_cong) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1555 | qed auto | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1556 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1557 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1558 | lemma summable_on_scaleR_right [intro]: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1559 | fixes f :: \<open>'a \<Rightarrow> 'b :: real_normed_vector\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1560 | assumes "c \<noteq> 0 \<Longrightarrow> f summable_on A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1561 | shows "(\<lambda>x. c *\<^sub>R f x) summable_on A" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1562 | proof (cases \<open>c = 0\<close>) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1563 | case False | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1564 | then have "(*\<^sub>R) c \<circ> f summable_on A" | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1565 | using assms by (auto simp add: scaleR_right.additive_axioms summable_on_comm_additive) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1566 | then show ?thesis | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1567 | by (metis (mono_tags, lifting) comp_apply summable_on_cong) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1568 | qed auto | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1569 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1570 | lemma infsum_scaleR_left: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1571 | fixes c :: \<open>'a :: real_normed_vector\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1572 | assumes "c \<noteq> 0 \<Longrightarrow> f summable_on A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1573 | shows "infsum (\<lambda>x. f x *\<^sub>R c) A = infsum f A *\<^sub>R c" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1574 | proof (cases \<open>c = 0\<close>) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1575 | case False | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1576 | then have "infsum ((\<lambda>y. y *\<^sub>R c) \<circ> f) A = infsum f A *\<^sub>R c" | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1577 | using assms by (auto simp add: scaleR_left.additive_axioms infsum_comm_additive) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1578 | then show ?thesis | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1579 | by (metis (mono_tags, lifting) comp_apply infsum_cong) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1580 | qed auto | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1581 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1582 | lemma infsum_scaleR_right: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1583 | fixes f :: \<open>'a \<Rightarrow> 'b :: real_normed_vector\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1584 | shows "infsum (\<lambda>x. c *\<^sub>R f x) A = c *\<^sub>R infsum f A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1585 | proof - | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1586 | consider (summable) \<open>f summable_on A\<close> | (c0) \<open>c = 0\<close> | (not_summable) \<open>\<not> f summable_on A\<close> \<open>c \<noteq> 0\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1587 | by auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1588 | then show ?thesis | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1589 | proof cases | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1590 | case summable | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1591 | then have "infsum ((*\<^sub>R) c \<circ> f) A = c *\<^sub>R infsum f A" | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1592 | by (auto simp add: scaleR_right.additive_axioms infsum_comm_additive) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1593 | then show ?thesis | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1594 | by (metis (mono_tags, lifting) comp_apply infsum_cong) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1595 | next | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1596 | case c0 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1597 | then show ?thesis by auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1598 | next | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1599 | case not_summable | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1600 | have \<open>\<not> (\<lambda>x. c *\<^sub>R f x) summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1601 | proof (rule notI) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1602 | assume \<open>(\<lambda>x. c *\<^sub>R f x) summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1603 | then have \<open>(\<lambda>x. inverse c *\<^sub>R c *\<^sub>R f x) summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1604 | using summable_on_scaleR_right by blast | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1605 | with not_summable show False | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1606 | by simp | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1607 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1608 | then show ?thesis | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1609 | by (simp add: infsum_not_exists not_summable(1)) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1610 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1611 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1612 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1613 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1614 | lemma infsum_Un_Int: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1615 |   fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add, t2_space}"
 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1616 | assumes "f summable_on A - B" "f summable_on B - A" \<open>f summable_on A \<inter> B\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1617 | shows "infsum f (A \<union> B) = infsum f A + infsum f B - infsum f (A \<inter> B)" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1618 | proof - | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1619 | obtain \<open>f summable_on A\<close> \<open>f summable_on B\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1620 | using assms by (metis Int_Diff_Un Int_Diff_disjoint inf_commute summable_on_Un_disjoint) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1621 | then have \<open>infsum f (A \<union> B) = infsum f A + infsum f (B - A)\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1622 | using assms(2) infsum_Un_disjoint by fastforce | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1623 | moreover have \<open>infsum f (B - A) = infsum f B - infsum f (A \<inter> B)\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1624 | using assms by (metis Diff_Int2 Un_Int_eq(2) \<open>f summable_on B\<close> inf_le2 infsum_Diff) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1625 | ultimately show ?thesis | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1626 | by auto | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1627 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1628 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1629 | lemma inj_combinator': | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1630 | assumes "x \<notin> F" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1631 | shows \<open>inj_on (\<lambda>(g, y). g(x := y)) (Pi\<^sub>E F B \<times> B x)\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1632 | proof - | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1633 | have "inj_on ((\<lambda>(y, g). g(x := y)) \<circ> prod.swap) (Pi\<^sub>E F B \<times> B x)" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1634 | using inj_combinator[of x F B] assms by (intro comp_inj_on) (auto simp: product_swap) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1635 | thus ?thesis | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1636 | by (simp add: o_def) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1637 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1638 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1639 | lemma infsum_prod_PiE: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1640 | \<comment> \<open>See also \<open>infsum_prod_PiE_abs\<close> below with incomparable premises.\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1641 |   fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c :: {comm_monoid_mult, topological_semigroup_mult, division_ring, banach}"
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1642 | assumes finite: "finite A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1643 | assumes "\<And>x. x \<in> A \<Longrightarrow> f x summable_on B x" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1644 | assumes "(\<lambda>g. \<Prod>x\<in>A. f x (g x)) summable_on (PiE A B)" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1645 | shows "infsum (\<lambda>g. \<Prod>x\<in>A. f x (g x)) (PiE A B) = (\<Prod>x\<in>A. infsum (f x) (B x))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1646 | proof (use finite assms(2-) in induction) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1647 | case empty | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1648 | then show ?case | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1649 | by auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1650 | next | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1651 | case (insert x F) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1652 | have pi: \<open>Pi\<^sub>E (insert x F) B = (\<lambda>(g,y). g(x:=y)) ` (Pi\<^sub>E F B \<times> B x)\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1653 | unfolding PiE_insert_eq | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1654 | by (subst swap_product [symmetric]) (simp add: image_image case_prod_unfold) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1655 | have prod: \<open>(\<Prod>x'\<in>F. f x' ((p(x:=y)) x')) = (\<Prod>x'\<in>F. f x' (p x'))\<close> for p y | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1656 | by (rule prod.cong) (use insert.hyps in auto) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1657 | have inj: \<open>inj_on (\<lambda>(g, y). g(x := y)) (Pi\<^sub>E F B \<times> B x)\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1658 | using \<open>x \<notin> F\<close> by (rule inj_combinator') | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1659 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1660 | have summable1: \<open>(\<lambda>g. \<Prod>x\<in>insert x F. f x (g x)) summable_on Pi\<^sub>E (insert x F) B\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1661 | using insert.prems(2) . | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1662 | also have \<open>Pi\<^sub>E (insert x F) B = (\<lambda>(g,y). g(x:=y)) ` (Pi\<^sub>E F B \<times> B x)\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1663 | by (simp only: pi) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1664 | also have "(\<lambda>g. \<Prod>x\<in>insert x F. f x (g x)) summable_on \<dots> \<longleftrightarrow> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1665 | ((\<lambda>g. \<Prod>x\<in>insert x F. f x (g x)) \<circ> (\<lambda>(g,y). g(x:=y))) summable_on (Pi\<^sub>E F B \<times> B x)" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1666 | using inj by (rule summable_on_reindex) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1667 | also have "(\<Prod>z\<in>F. f z ((g(x := y)) z)) = (\<Prod>z\<in>F. f z (g z))" for g y | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1668 | using insert.hyps by (intro prod.cong) auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1669 | hence "((\<lambda>g. \<Prod>x\<in>insert x F. f x (g x)) \<circ> (\<lambda>(g,y). g(x:=y))) = | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1670 | (\<lambda>(p, y). f x y * (\<Prod>x'\<in>F. f x' (p x')))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1671 | using insert.hyps by (auto simp: fun_eq_iff cong: prod.cong_simp) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1672 | finally have summable2: \<open>(\<lambda>(p, y). f x y * (\<Prod>x'\<in>F. f x' (p x'))) summable_on Pi\<^sub>E F B \<times> B x\<close> . | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1673 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1674 | then have \<open>(\<lambda>p. \<Sum>\<^sub>\<infinity>y\<in>B x. f x y * (\<Prod>x'\<in>F. f x' (p x'))) summable_on Pi\<^sub>E F B\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1675 | by (rule summable_on_Sigma_banach) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1676 | then have \<open>(\<lambda>p. (\<Sum>\<^sub>\<infinity>y\<in>B x. f x y) * (\<Prod>x'\<in>F. f x' (p x'))) summable_on Pi\<^sub>E F B\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1677 | by (metis (mono_tags, lifting) infsum_cmult_left' infsum_cong summable_on_cong) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1678 | then have summable3: \<open>(\<lambda>p. (\<Prod>x'\<in>F. f x' (p x'))) summable_on Pi\<^sub>E F B\<close> if \<open>(\<Sum>\<^sub>\<infinity>y\<in>B x. f x y) \<noteq> 0\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1679 | using summable_on_cmult_right' that by blast | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1680 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1681 | have \<open>(\<Sum>\<^sub>\<infinity>g\<in>Pi\<^sub>E (insert x F) B. \<Prod>x\<in>insert x F. f x (g x)) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1682 | = (\<Sum>\<^sub>\<infinity>(p,y)\<in>Pi\<^sub>E F B \<times> B x. \<Prod>x'\<in>insert x F. f x' ((p(x:=y)) x'))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1683 | by (smt (verit, ccfv_SIG) comp_apply infsum_cong infsum_reindex inj pi prod.cong split_def) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1684 | also have \<open>\<dots> = (\<Sum>\<^sub>\<infinity>(p, y)\<in>Pi\<^sub>E F B \<times> B x. f x y * (\<Prod>x'\<in>F. f x' ((p(x:=y)) x')))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1685 | using insert.hyps by auto | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1686 | also have \<open>\<dots> = (\<Sum>\<^sub>\<infinity>(p, y)\<in>Pi\<^sub>E F B \<times> B x. f x y * (\<Prod>x'\<in>F. f x' (p x')))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1687 | using prod by presburger | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1688 | also have \<open>\<dots> = (\<Sum>\<^sub>\<infinity>p\<in>Pi\<^sub>E F B. \<Sum>\<^sub>\<infinity>y\<in>B x. f x y * (\<Prod>x'\<in>F. f x' (p x')))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1689 | using infsum_Sigma'_banach summable2 by force | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1690 | also have \<open>\<dots> = (\<Sum>\<^sub>\<infinity>y\<in>B x. f x y) * (\<Sum>\<^sub>\<infinity>p\<in>Pi\<^sub>E F B. \<Prod>x'\<in>F. f x' (p x'))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1691 | by (smt (verit) infsum_cmult_left' infsum_cmult_right' infsum_cong) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1692 | also have \<open>\<dots> = (\<Prod>x\<in>insert x F. infsum (f x) (B x))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1693 | using insert summable3 by auto | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1694 | finally show ?case | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1695 | by simp | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1696 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1697 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1698 | lemma infsum_prod_PiE_abs: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1699 |   \<comment> \<open>See also @{thm [source] infsum_prod_PiE} above with incomparable premises.\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1700 |   fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c :: {banach, real_normed_div_algebra, comm_semiring_1}"
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1701 | assumes finite: "finite A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1702 | assumes "\<And>x. x \<in> A \<Longrightarrow> f x abs_summable_on B x" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1703 | shows "infsum (\<lambda>g. \<Prod>x\<in>A. f x (g x)) (PiE A B) = (\<Prod>x\<in>A. infsum (f x) (B x))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1704 | proof (use finite assms(2) in induction) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1705 | case empty | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1706 | then show ?case | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1707 | by auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1708 | next | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1709 | case (insert x A) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1710 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1711 | have pi: \<open>Pi\<^sub>E (insert x F) B = (\<lambda>(g,y). g(x:=y)) ` (Pi\<^sub>E F B \<times> B x)\<close> for x F and B :: "'a \<Rightarrow> 'b set" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1712 | unfolding PiE_insert_eq | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1713 | by (subst swap_product [symmetric]) (simp add: image_image case_prod_unfold) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1714 | have prod: \<open>(\<Prod>x'\<in>A. f x' ((p(x:=y)) x')) = (\<Prod>x'\<in>A. f x' (p x'))\<close> for p y | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1715 | by (rule prod.cong) (use insert.hyps in auto) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1716 | have inj: \<open>inj_on (\<lambda>(g, y). g(x := y)) (Pi\<^sub>E A B \<times> B x)\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1717 | using \<open>x \<notin> A\<close> by (rule inj_combinator') | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1718 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1719 | define s where \<open>s x = infsum (\<lambda>y. norm (f x y)) (B x)\<close> for x | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1720 | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1721 | have \<open>(\<Sum>p\<in>P. norm (\<Prod>x\<in>F. f x (p x))) \<le> prod s F\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1722 | if P: \<open>P \<subseteq> Pi\<^sub>E F B\<close> and [simp]: \<open>finite P\<close> \<open>finite F\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1723 | and sum: \<open>\<And>x. x \<in> F \<Longrightarrow> f x abs_summable_on B x\<close> for P F | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1724 | proof - | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1725 |     define B' where \<open>B' x = {p x| p. p\<in>P}\<close> for x
 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1726 | have fin_B'[simp]: \<open>finite (B' x)\<close> for x | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1727 | using that by (auto simp: B'_def) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1728 | have [simp]: \<open>finite (Pi\<^sub>E F B')\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1729 | by (simp add: finite_PiE) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1730 | have [simp]: \<open>P \<subseteq> Pi\<^sub>E F B'\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1731 | using that by (auto simp: B'_def) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1732 | have B'B: \<open>B' x \<subseteq> B x\<close> if \<open>x \<in> F\<close> for x | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1733 | unfolding B'_def using P that | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1734 | by auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1735 | have s_bound: \<open>(\<Sum>y\<in>B' x. norm (f x y)) \<le> s x\<close> if \<open>x \<in> F\<close> for x | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1736 | by (metis B'B fin_B' finite_sum_le_has_sum has_sum_infsum norm_ge_zero s_def sum that) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1737 | have \<open>(\<Sum>p\<in>P. norm (\<Prod>x\<in>F. f x (p x))) \<le> (\<Sum>p\<in>Pi\<^sub>E F B'. norm (\<Prod>x\<in>F. f x (p x)))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1738 | by (simp add: sum_mono2) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1739 | also have \<open>\<dots> = (\<Sum>p\<in>Pi\<^sub>E F B'. \<Prod>x\<in>F. norm (f x (p x)))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1740 | by (simp add: prod_norm) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1741 | also have \<open>\<dots> = (\<Prod>x\<in>F. \<Sum>y\<in>B' x. norm (f x y))\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1742 | proof (use \<open>finite F\<close> in induction) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1743 | case empty | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1744 | then show ?case by simp | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1745 | next | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1746 | case (insert x F) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1747 | have inj: \<open>inj_on (\<lambda>(g, y). g(x := y)) (Pi\<^sub>E F B' \<times> B' x)\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1748 | by (simp add: inj_combinator' insert.hyps) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1749 | then have \<open>(\<Sum>p\<in>Pi\<^sub>E (insert x F) B'. \<Prod>x\<in>insert x F. norm (f x (p x))) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1750 | = (\<Sum>(p,y)\<in>Pi\<^sub>E F B' \<times> B' x. \<Prod>x'\<in>insert x F. norm (f x' ((p(x := y)) x')))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1751 | by (simp add: pi sum.reindex case_prod_unfold) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1752 | also have \<open>\<dots> = (\<Sum>(p, y)\<in>Pi\<^sub>E F B' \<times> B' x. norm (f x y) * (\<Prod>x'\<in>F. norm (f x' (p x'))))\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1753 | by (smt (verit, del_insts) fun_upd_apply insert.hyps prod.cong prod.insert split_def sum.cong) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1754 | also have \<open>\<dots> = (\<Sum>y\<in>B' x. norm (f x y)) * (\<Sum>p\<in>Pi\<^sub>E F B'. \<Prod>x'\<in>F. norm (f x' (p x')))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1755 | by (simp add: sum_product sum.swap [of _ "Pi\<^sub>E F B'"] sum.cartesian_product) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1756 | also have \<open>\<dots> = (\<Prod>x\<in>insert x F. \<Sum>y\<in>B' x. norm (f x y))\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1757 | using insert by force | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1758 | finally show ?case . | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1759 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1760 | also have \<open>\<dots> \<le> (\<Prod>x\<in>F. s x)\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1761 | using s_bound by (simp add: prod_mono sum_nonneg) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1762 | finally show ?thesis . | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1763 | qed | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1764 | then have "bdd_above | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1765 |      (sum (\<lambda>g. norm (\<Prod>x\<in>insert x A. f x (g x))) ` {F. F \<subseteq> Pi\<^sub>E (insert x A) B \<and> finite F})"
 | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1766 | using insert.hyps insert.prems by (intro bdd_aboveI) blast | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1767 | then have \<open>(\<lambda>g. \<Prod>x\<in>insert x A. f x (g x)) abs_summable_on Pi\<^sub>E (insert x A) B\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1768 | using nonneg_bdd_above_summable_on | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1769 | by (metis (mono_tags, lifting) Collect_cong norm_ge_zero) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1770 | also have \<open>Pi\<^sub>E (insert x A) B = (\<lambda>(g,y). g(x:=y)) ` (Pi\<^sub>E A B \<times> B x)\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1771 | by (simp only: pi) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1772 | also have "(\<lambda>g. \<Prod>x\<in>insert x A. f x (g x)) abs_summable_on \<dots> \<longleftrightarrow> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1773 | ((\<lambda>g. \<Prod>x\<in>insert x A. f x (g x)) \<circ> (\<lambda>(g,y). g(x:=y))) abs_summable_on (Pi\<^sub>E A B \<times> B x)" | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1774 | using inj by (subst summable_on_reindex) (auto simp: o_def) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1775 | also have "(\<Prod>z\<in>A. f z ((g(x := y)) z)) = (\<Prod>z\<in>A. f z (g z))" for g y | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1776 | using insert.hyps by (intro prod.cong) auto | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1777 | hence "((\<lambda>g. \<Prod>x\<in>insert x A. f x (g x)) \<circ> (\<lambda>(g,y). g(x:=y))) = | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1778 | (\<lambda>(p, y). f x y * (\<Prod>x'\<in>A. f x' (p x')))" | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1779 | using insert.hyps by (auto simp: fun_eq_iff cong: prod.cong_simp) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1780 | finally have summable2: \<open>(\<lambda>(p, y). f x y * (\<Prod>x'\<in>A. f x' (p x'))) abs_summable_on Pi\<^sub>E A B \<times> B x\<close> . | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1781 | |
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1782 | have \<open>(\<Sum>\<^sub>\<infinity>g\<in>Pi\<^sub>E (insert x A) B. \<Prod>x\<in>insert x A. f x (g x)) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1783 | = (\<Sum>\<^sub>\<infinity>(p,y)\<in>Pi\<^sub>E A B \<times> B x. \<Prod>x'\<in>insert x A. f x' ((p(x:=y)) x'))\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1784 | using inj by (simp add: pi infsum_reindex o_def case_prod_unfold) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1785 | also have \<open>\<dots> = (\<Sum>\<^sub>\<infinity>(p,y) \<in> Pi\<^sub>E A B \<times> B x. f x y * (\<Prod>x'\<in>A. f x' (p x')))\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1786 | using prod insert.hyps by auto | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1787 | also have \<open>\<dots> = (\<Sum>\<^sub>\<infinity>p\<in>Pi\<^sub>E A B. \<Sum>\<^sub>\<infinity>y\<in>B x. f x y * (\<Prod>x'\<in>A. f x' (p x')))\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1788 | using abs_summable_summable infsum_Sigma'_banach summable2 by fastforce | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1789 | also have \<open>\<dots> = (\<Sum>\<^sub>\<infinity>y\<in>B x. f x y) * (\<Sum>\<^sub>\<infinity>p\<in>Pi\<^sub>E A B. \<Prod>x'\<in>A. f x' (p x'))\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1790 | by (smt (verit, best) infsum_cmult_left' infsum_cmult_right' infsum_cong) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1791 | finally show ?case | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1792 | by (simp add: insert) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1793 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1794 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1795 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1796 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1797 | subsection \<open>Absolute convergence\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1798 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1799 | lemma abs_summable_countable: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1800 | assumes \<open>f abs_summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1801 |   shows \<open>countable {x\<in>A. f x \<noteq> 0}\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1802 | proof - | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1803 |   have fin: \<open>finite {x\<in>A. norm (f x) \<ge> t}\<close> if \<open>t > 0\<close> for t
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1804 | proof (rule ccontr) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1805 |     assume *: \<open>infinite {x \<in> A. t \<le> norm (f x)}\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1806 | have \<open>infsum (\<lambda>x. norm (f x)) A \<ge> b\<close> for b | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1807 | proof - | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1808 | obtain b' where b': \<open>of_nat b' \<ge> b / t\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1809 | by (meson real_arch_simple) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1810 | from * | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1811 |       obtain F where cardF: \<open>card F \<ge> b'\<close> and \<open>finite F\<close> and F: \<open>F \<subseteq> {x \<in> A. t \<le> norm (f x)}\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1812 | by (meson finite_if_finite_subsets_card_bdd nle_le) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1813 | have \<open>b \<le> of_nat b' * t\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1814 | using b' \<open>t > 0\<close> by (simp add: field_simps split: if_splits) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1815 | also have \<open>\<dots> \<le> of_nat (card F) * t\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1816 | by (simp add: cardF that) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1817 | also have \<open>\<dots> = sum (\<lambda>x. t) F\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1818 | by simp | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1819 | also have \<open>\<dots> \<le> sum (\<lambda>x. norm (f x)) F\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1820 | by (metis (mono_tags, lifting) F in_mono mem_Collect_eq sum_mono) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1821 | also have \<open>\<dots> = infsum (\<lambda>x. norm (f x)) F\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1822 | using \<open>finite F\<close> by (rule infsum_finite[symmetric]) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1823 | also have \<open>\<dots> \<le> infsum (\<lambda>x. norm (f x)) A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1824 | by (rule infsum_mono_neutral) (use \<open>finite F\<close> assms F in auto) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1825 | finally show ?thesis . | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1826 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1827 | then show False | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1828 | by (meson gt_ex linorder_not_less) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1829 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1830 |   have \<open>countable (\<Union>i\<in>{1..}. {x\<in>A. norm (f x) \<ge> 1/of_nat i})\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1831 | by (rule countable_UN) (use fin in \<open>auto intro!: countable_finite\<close>) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1832 |   also have \<open>\<dots> = {x\<in>A. f x \<noteq> 0}\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1833 | proof safe | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1834 | fix x assume x: "x \<in> A" "f x \<noteq> 0" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1835 | define i where "i = max 1 (nat (ceiling (1 / norm (f x))))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1836 | have "i \<ge> 1" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1837 | by (simp add: i_def) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1838 | moreover have "real i \<ge> 1 / norm (f x)" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1839 | unfolding i_def by linarith | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1840 | hence "1 / real i \<le> norm (f x)" using \<open>f x \<noteq> 0\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1841 | by (auto simp: divide_simps mult_ac) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1842 |     ultimately show "x \<in> (\<Union>i\<in>{1..}. {x \<in> A. 1 / real i \<le> norm (f x)})"
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1843 | using \<open>x \<in> A\<close> by auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1844 | qed auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1845 | finally show ?thesis . | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1846 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1847 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1848 | (* Logically belongs in the section about reals, but needed as a dependency here *) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1849 | lemma summable_on_iff_abs_summable_on_real: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1850 | fixes f :: \<open>'a \<Rightarrow> real\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1851 | shows \<open>f summable_on A \<longleftrightarrow> f abs_summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1852 | proof (rule iffI) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1853 | assume \<open>f summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1854 | define n A\<^sub>p A\<^sub>n | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1855 |     where \<open>n \<equiv> \<lambda>x. norm (f x)\<close> and \<open>A\<^sub>p = {x\<in>A. f x \<ge> 0}\<close> and \<open>A\<^sub>n = {x\<in>A. f x < 0}\<close> for x
 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1856 |   have A: \<open>A\<^sub>p \<union> A\<^sub>n = A\<close> \<open>A\<^sub>p \<inter> A\<^sub>n = {}\<close>
 | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1857 | by (auto simp: A\<^sub>p_def A\<^sub>n_def) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1858 | from \<open>f summable_on A\<close> have \<open>f summable_on A\<^sub>p\<close> \<open>f summable_on A\<^sub>n\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1859 | using A\<^sub>p_def A\<^sub>n_def summable_on_subset_banach by fastforce+ | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1860 | then have \<open>n summable_on A\<^sub>p\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1861 | by (smt (verit) A\<^sub>p_def n_def mem_Collect_eq real_norm_def summable_on_cong) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1862 | moreover have \<open>n summable_on A\<^sub>n\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1863 | by (smt (verit, best) \<open>f summable_on A\<^sub>n\<close> summable_on_uminus A\<^sub>n_def n_def summable_on_cong mem_Collect_eq real_norm_def) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1864 | ultimately show \<open>n summable_on A\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1865 | using A summable_on_Un_disjoint by blast | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1866 | next | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1867 | show \<open>f abs_summable_on A \<Longrightarrow> f summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1868 | using abs_summable_summable by blast | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1869 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1870 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1871 | lemma abs_summable_on_Sigma_iff: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1872 | shows "f abs_summable_on Sigma A B \<longleftrightarrow> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1873 | (\<forall>x\<in>A. (\<lambda>y. f (x, y)) abs_summable_on B x) \<and> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1874 | ((\<lambda>x. infsum (\<lambda>y. norm (f (x, y))) (B x)) abs_summable_on A)" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1875 | proof (intro iffI conjI ballI) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1876 | assume asm: \<open>f abs_summable_on Sigma A B\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1877 | then have \<open>(\<lambda>x. infsum (\<lambda>y. norm (f (x,y))) (B x)) summable_on A\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1878 | by (simp add: cond_case_prod_eta summable_on_Sigma_banach) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1879 | then show \<open>(\<lambda>x. \<Sum>\<^sub>\<infinity>y\<in>B x. norm (f (x, y))) abs_summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1880 | using summable_on_iff_abs_summable_on_real by force | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1881 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1882 | show \<open>(\<lambda>y. f (x, y)) abs_summable_on B x\<close> if \<open>x \<in> A\<close> for x | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1883 | proof - | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1884 | from asm have \<open>f abs_summable_on Pair x ` B x\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1885 | by (simp add: image_subset_iff summable_on_subset_banach that) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1886 | then show ?thesis | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1887 | by (metis (mono_tags, lifting) o_def inj_on_def summable_on_reindex prod.inject summable_on_cong) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1888 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1889 | next | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1890 | assume asm: \<open>(\<forall>x\<in>A. (\<lambda>xa. f (x, xa)) abs_summable_on B x) \<and> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1891 | (\<lambda>x. \<Sum>\<^sub>\<infinity>y\<in>B x. norm (f (x, y))) abs_summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1892 | have \<open>(\<Sum>xy\<in>F. norm (f xy)) \<le> (\<Sum>\<^sub>\<infinity>x\<in>A. \<Sum>\<^sub>\<infinity>y\<in>B x. norm (f (x, y)))\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1893 | if \<open>F \<subseteq> Sigma A B\<close> and [simp]: \<open>finite F\<close> for F | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1894 | proof - | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1895 |     have [simp]: \<open>(SIGMA x:fst ` F. {y. (x, y) \<in> F}) = F\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1896 | by (auto intro!: set_eqI simp add: Domain.DomainI fst_eq_Domain) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1897 |     have [simp]: \<open>finite {y. (x, y) \<in> F}\<close> for x
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1898 | by (metis \<open>finite F\<close> Range.intros finite_Range finite_subset mem_Collect_eq subsetI) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1899 |     have \<open>(\<Sum>xy\<in>F. norm (f xy)) = (\<Sum>x\<in>fst ` F. \<Sum>y\<in>{y. (x,y)\<in>F}. norm (f (x,y)))\<close>
 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1900 | by (simp add: sum.Sigma) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1901 |     also have \<open>\<dots> = (\<Sum>\<^sub>\<infinity>x\<in>fst ` F. \<Sum>\<^sub>\<infinity>y\<in>{y. (x,y)\<in>F}. norm (f (x,y)))\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1902 | by auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1903 | also have \<open>\<dots> \<le> (\<Sum>\<^sub>\<infinity>x\<in>fst ` F. \<Sum>\<^sub>\<infinity>y\<in>B x. norm (f (x,y)))\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1904 | using asm that(1) by (intro infsum_mono infsum_mono_neutral) auto | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1905 | also have \<open>\<dots> \<le> (\<Sum>\<^sub>\<infinity>x\<in>A. \<Sum>\<^sub>\<infinity>y\<in>B x. norm (f (x,y)))\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1906 | by (rule infsum_mono_neutral) (use asm that(1) in \<open>auto simp add: infsum_nonneg\<close>) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1907 | finally show ?thesis . | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1908 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1909 | then show \<open>f abs_summable_on Sigma A B\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1910 | by (intro nonneg_bdd_above_summable_on) (auto simp: bdd_above_def) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1911 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1912 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1913 | lemma abs_summable_on_comparison_test: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1914 | assumes "g abs_summable_on A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1915 | assumes "\<And>x. x \<in> A \<Longrightarrow> norm (f x) \<le> norm (g x)" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1916 | shows "f abs_summable_on A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1917 | proof (rule nonneg_bdd_above_summable_on) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1918 |   show "bdd_above (sum (\<lambda>x. norm (f x)) ` {F. F \<subseteq> A \<and> finite F})"
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1919 | proof (rule bdd_aboveI2) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1920 |     fix F assume F: "F \<in> {F. F \<subseteq> A \<and> finite F}"
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1921 | have \<open>sum (\<lambda>x. norm (f x)) F \<le> sum (\<lambda>x. norm (g x)) F\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1922 | using assms F by (intro sum_mono) auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1923 | also have \<open>\<dots> = infsum (\<lambda>x. norm (g x)) F\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1924 | using F by simp | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1925 | also have \<open>\<dots> \<le> infsum (\<lambda>x. norm (g x)) A\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1926 | by (smt (verit) F assms(1) infsum_mono2 mem_Collect_eq norm_ge_zero summable_on_subset_banach) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1927 | finally show "(\<Sum>x\<in>F. norm (f x)) \<le> (\<Sum>\<^sub>\<infinity>x\<in>A. norm (g x))" . | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1928 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1929 | qed auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1930 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1931 | lemma abs_summable_iff_bdd_above: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1932 | fixes f :: \<open>'a \<Rightarrow> 'b::real_normed_vector\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1933 |   shows \<open>f abs_summable_on A \<longleftrightarrow> bdd_above (sum (\<lambda>x. norm (f x)) ` {F. F\<subseteq>A \<and> finite F})\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1934 | proof (rule iffI) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1935 | assume \<open>f abs_summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1936 |   show \<open>bdd_above (sum (\<lambda>x. norm (f x)) ` {F. F \<subseteq> A \<and> finite F})\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1937 | proof (rule bdd_aboveI2) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1938 |     fix F assume F: "F \<in> {F. F \<subseteq> A \<and> finite F}"
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1939 | show "(\<Sum>x\<in>F. norm (f x)) \<le> (\<Sum>\<^sub>\<infinity>x\<in>A. norm (f x))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1940 | by (rule finite_sum_le_infsum) (use \<open>f abs_summable_on A\<close> F in auto) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1941 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1942 | next | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1943 |   assume \<open>bdd_above (sum (\<lambda>x. norm (f x)) ` {F. F\<subseteq>A \<and> finite F})\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1944 | then show \<open>f abs_summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1945 | by (simp add: nonneg_bdd_above_summable_on) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1946 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1947 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1948 | lemma abs_summable_product: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1949 |   fixes x :: "'a \<Rightarrow> 'b::{real_normed_div_algebra,banach,second_countable_topology}"
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1950 | assumes x2_sum: "(\<lambda>i. (x i) * (x i)) abs_summable_on A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1951 | and y2_sum: "(\<lambda>i. (y i) * (y i)) abs_summable_on A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1952 | shows "(\<lambda>i. x i * y i) abs_summable_on A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1953 | proof (rule nonneg_bdd_above_summable_on) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1954 |   show "bdd_above (sum (\<lambda>xa. norm (x xa * y xa)) ` {F. F \<subseteq> A \<and> finite F})"
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1955 | proof (rule bdd_aboveI2) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1956 |     fix F assume F: \<open>F \<in> {F. F \<subseteq> A \<and> finite F}\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1957 | then have r1: "finite F" and b4: "F \<subseteq> A" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1958 | by auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1959 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1960 | have a1: "(\<Sum>\<^sub>\<infinity>i\<in>F. norm (x i * x i)) \<le> (\<Sum>\<^sub>\<infinity>i\<in>A. norm (x i * x i))" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 1961 | by (metis (no_types, lifting) b4 infsum_mono2 norm_ge_zero summable_on_subset_banach x2_sum) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1962 | |
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1963 | have "norm (x i * y i) \<le> norm (x i * x i) + norm (y i * y i)" for i | 
| 79757 
f20ac6788faa
tuned proof: avoid z3 to make it work on arm64-linux;
 wenzelm parents: 
79529diff
changeset | 1964 | unfolding norm_mult by (smt (verit, best) abs_norm_cancel mult_mono not_sum_squares_lt_zero) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1965 | hence "(\<Sum>i\<in>F. norm (x i * y i)) \<le> (\<Sum>i\<in>F. norm (x i * x i) + norm (y i * y i))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1966 | by (simp add: sum_mono) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1967 | also have "\<dots> = (\<Sum>i\<in>F. norm (x i * x i)) + (\<Sum>i\<in>F. norm (y i * y i))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1968 | by (simp add: sum.distrib) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1969 | also have "\<dots> = (\<Sum>\<^sub>\<infinity>i\<in>F. norm (x i * x i)) + (\<Sum>\<^sub>\<infinity>i\<in>F. norm (y i * y i))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1970 | by (simp add: \<open>finite F\<close>) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1971 | also have "\<dots> \<le> (\<Sum>\<^sub>\<infinity>i\<in>A. norm (x i * x i)) + (\<Sum>\<^sub>\<infinity>i\<in>A. norm (y i * y i))" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1972 | using F assms | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1973 | by (intro add_mono infsum_mono2) auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1974 | finally show \<open>(\<Sum>xa\<in>F. norm (x xa * y xa)) \<le> (\<Sum>\<^sub>\<infinity>i\<in>A. norm (x i * x i)) + (\<Sum>\<^sub>\<infinity>i\<in>A. norm (y i * y i))\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1975 | by simp | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1976 | qed | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1977 | qed auto | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 1978 | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1979 | subsection \<open>Extended reals and nats\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1980 | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1981 | lemma summable_on_ennreal[simp]: \<open>(f::_ \<Rightarrow> ennreal) summable_on S\<close> and summable_on_enat[simp]: \<open>(f::_ \<Rightarrow> enat) summable_on S\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1982 | by (simp_all add: nonneg_summable_on_complete) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1983 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1984 | lemma has_sum_superconst_infinite_ennreal: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1985 | fixes f :: \<open>'a \<Rightarrow> ennreal\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1986 | assumes geqb: \<open>\<And>x. x \<in> S \<Longrightarrow> f x \<ge> b\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1987 | assumes b: \<open>b > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1988 | assumes \<open>infinite S\<close> | 
| 77359 | 1989 | shows "(f has_sum \<infinity>) S" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1990 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1991 | have \<open>(sum f \<longlongrightarrow> \<infinity>) (finite_subsets_at_top S)\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1992 | proof (rule order_tendstoI) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1993 | fix y :: ennreal assume \<open>y < \<infinity>\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1994 | then have \<open>y / b < \<infinity>\<close> \<open>y < top\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 1995 | using b ennreal_divide_eq_top_iff top.not_eq_extremum by force+ | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1996 | then obtain F where \<open>finite F\<close> and \<open>F \<subseteq> S\<close> and cardF: \<open>card F > y / b\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1997 | using \<open>infinite S\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1998 | by (metis ennreal_Ex_less_of_nat infinite_arbitrarily_large infinity_ennreal_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1999 | moreover have \<open>sum f Y > y\<close> if \<open>finite Y\<close> and \<open>F \<subseteq> Y\<close> and \<open>Y \<subseteq> S\<close> for Y | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2000 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2001 | have \<open>y < b * card F\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2002 | by (metis b \<open>y < top\<close> cardF divide_less_ennreal ennreal_mult_eq_top_iff gr_implies_not_zero mult.commute top.not_eq_extremum) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2003 | also have \<open>\<dots> \<le> b * card Y\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2004 | by (meson b card_mono less_imp_le mult_left_mono of_nat_le_iff that) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2005 | also have \<open>\<dots> = sum (\<lambda>_. b) Y\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2006 | by (simp add: mult.commute) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2007 | also have \<open>\<dots> \<le> sum f Y\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2008 | using geqb by (meson subset_eq sum_mono that(3)) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2009 | finally show ?thesis . | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2010 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2011 | ultimately show \<open>\<forall>\<^sub>F x in finite_subsets_at_top S. y < sum f x\<close> | 
| 77359 | 2012 | unfolding eventually_finite_subsets_at_top by auto | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2013 | qed auto | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2014 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2015 | by (simp add: has_sum_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2016 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2017 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2018 | lemma infsum_superconst_infinite_ennreal: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2019 | fixes f :: \<open>'a \<Rightarrow> ennreal\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2020 | assumes \<open>\<And>x. x \<in> S \<Longrightarrow> f x \<ge> b\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2021 | assumes \<open>b > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2022 | assumes \<open>infinite S\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2023 | shows "infsum f S = \<infinity>" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2024 | using assms infsumI has_sum_superconst_infinite_ennreal by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2025 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2026 | lemma infsum_superconst_infinite_ereal: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2027 | fixes f :: \<open>'a \<Rightarrow> ereal\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2028 | assumes geqb: \<open>\<And>x. x \<in> S \<Longrightarrow> f x \<ge> b\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2029 | assumes b: \<open>b > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2030 | assumes \<open>infinite S\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2031 | shows "infsum f S = \<infinity>" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2032 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2033 | obtain b' where b': \<open>e2ennreal b' = b\<close> and \<open>b' > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2034 | using b by blast | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2035 | have "0 < e2ennreal b" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2036 | using b' b | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2037 | by (metis dual_order.refl enn2ereal_e2ennreal gr_zeroI order_less_le zero_ennreal.abs_eq) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2038 | hence *: \<open>infsum (e2ennreal \<circ> f) S = \<infinity>\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2039 | using assms b' | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2040 | by (intro infsum_superconst_infinite_ennreal[where b=b']) (auto intro!: e2ennreal_mono) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2041 | have \<open>infsum f S = infsum (enn2ereal \<circ> (e2ennreal \<circ> f)) S\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2042 | using geqb b by (intro infsum_cong) (fastforce simp: enn2ereal_e2ennreal) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2043 | also have \<open>\<dots> = enn2ereal \<infinity>\<close> | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2044 | using * by (simp add: infsum_comm_additive_general continuous_at_enn2ereal nonneg_summable_on_complete) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2045 | also have \<open>\<dots> = \<infinity>\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2046 | by simp | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2047 | finally show ?thesis . | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2048 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2049 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2050 | lemma has_sum_superconst_infinite_ereal: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2051 | fixes f :: \<open>'a \<Rightarrow> ereal\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2052 | assumes \<open>\<And>x. x \<in> S \<Longrightarrow> f x \<ge> b\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2053 | assumes \<open>b > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2054 | assumes \<open>infinite S\<close> | 
| 77359 | 2055 | shows "(f has_sum \<infinity>) S" | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2056 | by (metis Infty_neq_0(1) assms infsum_def has_sum_infsum infsum_superconst_infinite_ereal) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2057 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2058 | lemma infsum_superconst_infinite_enat: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2059 | fixes f :: \<open>'a \<Rightarrow> enat\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2060 | assumes geqb: \<open>\<And>x. x \<in> S \<Longrightarrow> f x \<ge> b\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2061 | assumes b: \<open>b > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2062 | assumes \<open>infinite S\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2063 | shows "infsum f S = \<infinity>" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2064 | proof - | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2065 | have \<open>ennreal_of_enat (infsum f S) = infsum (ennreal_of_enat \<circ> f) S\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2066 | by (simp flip: infsum_comm_additive_general) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2067 | also have \<open>\<dots> = \<infinity>\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2068 | by (metis assms(3) b comp_def ennreal_of_enat_0 ennreal_of_enat_le_iff geqb infsum_superconst_infinite_ennreal leD leI) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2069 | also have \<open>\<dots> = ennreal_of_enat \<infinity>\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2070 | by simp | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2071 | finally show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2072 | by (rule ennreal_of_enat_inj[THEN iffD1]) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2073 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2074 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2075 | lemma has_sum_superconst_infinite_enat: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2076 | fixes f :: \<open>'a \<Rightarrow> enat\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2077 | assumes \<open>\<And>x. x \<in> S \<Longrightarrow> f x \<ge> b\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2078 | assumes \<open>b > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2079 | assumes \<open>infinite S\<close> | 
| 77359 | 2080 | shows "(f has_sum \<infinity>) S" | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2081 | by (metis assms i0_lb has_sum_infsum infsum_superconst_infinite_enat nonneg_summable_on_complete) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2082 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2083 | text \<open>This lemma helps to relate a real-valued infsum to a supremum over extended nonnegative reals.\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2084 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2085 | lemma infsum_nonneg_is_SUPREMUM_ennreal: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2086 | fixes f :: "'a \<Rightarrow> real" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2087 | assumes summable: "f summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2088 | and fnn: "\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2089 |   shows "ennreal (infsum f A) = (SUP F\<in>{F. finite F \<and> F \<subseteq> A}. (ennreal (sum f F)))"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2090 | proof - | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2091 | have \<section>: "\<And>F. \<lbrakk>finite F; F \<subseteq> A\<rbrakk> \<Longrightarrow> sum (ennreal \<circ> f) F = ennreal (sum f F)" | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2092 | by (metis (mono_tags, lifting) comp_def fnn subsetD sum.cong sum_ennreal) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2093 | then have \<open>ennreal (infsum f A) = infsum (ennreal \<circ> f) A\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2094 | by (simp add: infsum_comm_additive_general local.summable) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2095 |   also have \<open>\<dots> = (SUP F\<in>{F. finite F \<and> F \<subseteq> A}. (ennreal (sum f F)))\<close>
 | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2096 | by (metis (mono_tags, lifting) \<section> image_cong mem_Collect_eq nonneg_infsum_complete zero_le) | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2097 | finally show ?thesis . | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2098 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2099 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2100 | text \<open>This lemma helps to related a real-valued infsum to a supremum over extended reals.\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2101 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2102 | lemma infsum_nonneg_is_SUPREMUM_ereal: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2103 | fixes f :: "'a \<Rightarrow> real" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2104 | assumes summable: "f summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2105 | and fnn: "\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2106 |   shows "ereal (infsum f A) = (SUP F\<in>{F. finite F \<and> F \<subseteq> A}. (ereal (sum f F)))"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2107 | proof - | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2108 | have "\<And>F. \<lbrakk>finite F; F \<subseteq> A\<rbrakk> \<Longrightarrow> sum (ereal \<circ> f) F = ereal (sum f F)" | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2109 | by auto | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2110 | then have \<open>ereal (infsum f A) = infsum (ereal \<circ> f) A\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2111 | by (simp add: infsum_comm_additive_general local.summable) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2112 |   also have \<open>\<dots> = (SUP F\<in>{F. finite F \<and> F \<subseteq> A}. (ereal (sum f F)))\<close>
 | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2113 | by (subst nonneg_infsum_complete) (simp_all add: assms) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2114 | finally show ?thesis . | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2115 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2116 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2117 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2118 | subsection \<open>Real numbers\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2119 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2120 | text \<open>Most lemmas in the general property section already apply to real numbers. | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2121 | A few ones that are specific to reals are given here.\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2122 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2123 | lemma infsum_nonneg_is_SUPREMUM_real: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2124 | fixes f :: "'a \<Rightarrow> real" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2125 | assumes summable: "f summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2126 | and fnn: "\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2127 |   shows "infsum f A = (SUP F\<in>{F. finite F \<and> F \<subseteq> A}. (sum f F))"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2128 | proof - | 
| 77388 | 2129 |   have *: "ereal (infsum f A) = (SUP F\<in>{F. finite F \<and> F \<subseteq> A}. (ereal (sum f F)))"
 | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2130 | using assms by (rule infsum_nonneg_is_SUPREMUM_ereal) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2131 |   also have "\<dots> = ereal (SUP F\<in>{F. finite F \<and> F \<subseteq> A}. (sum f F))"
 | 
| 77388 | 2132 | by (metis (no_types, lifting) * MInfty_neq_ereal(2) PInfty_neq_ereal(2) SUP_cong abs_eq_infinity_cases ereal_SUP) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2133 | finally show ?thesis by simp | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2134 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2135 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2136 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2137 | lemma has_sum_nonneg_SUPREMUM_real: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2138 | fixes f :: "'a \<Rightarrow> real" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2139 | assumes "f summable_on A" and "\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0" | 
| 77359 | 2140 |   shows "(f has_sum (SUP F\<in>{F. finite F \<and> F \<subseteq> A}. (sum f F))) A"
 | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2141 | by (metis (mono_tags, lifting) assms has_sum_infsum infsum_nonneg_is_SUPREMUM_real) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2142 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2143 | lemma summable_countable_real: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2144 | fixes f :: \<open>'a \<Rightarrow> real\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2145 | assumes \<open>f summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2146 |   shows \<open>countable {x\<in>A. f x \<noteq> 0}\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2147 | using abs_summable_countable assms summable_on_iff_abs_summable_on_real by blast | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2148 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2149 | subsection \<open>Complex numbers\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2150 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2151 | lemma has_sum_cnj_iff[simp]: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2152 | fixes f :: \<open>'a \<Rightarrow> complex\<close> | 
| 77359 | 2153 | shows \<open>((\<lambda>x. cnj (f x)) has_sum cnj a) M \<longleftrightarrow> (f has_sum a) M\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2154 | by (simp add: has_sum_def lim_cnj del: cnj_sum add: cnj_sum[symmetric, abs_def, of f]) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2155 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2156 | lemma summable_on_cnj_iff[simp]: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2157 | "(\<lambda>i. cnj (f i)) summable_on A \<longleftrightarrow> f summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2158 | by (metis complex_cnj_cnj summable_on_def has_sum_cnj_iff) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2159 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2160 | lemma infsum_cnj[simp]: \<open>infsum (\<lambda>x. cnj (f x)) M = cnj (infsum f M)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2161 | by (metis complex_cnj_zero infsumI has_sum_cnj_iff infsum_def summable_on_cnj_iff has_sum_infsum) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2162 | |
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2163 | lemma has_sum_Re: | 
| 77359 | 2164 | assumes "(f has_sum a) M" | 
| 2165 | shows "((\<lambda>x. Re (f x)) has_sum Re a) M" | |
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2166 | using has_sum_comm_additive[where f=Re] | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2167 | using assms tendsto_Re by (fastforce simp add: o_def additive_def) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2168 | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2169 | lemma infsum_Re: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2170 | assumes "f summable_on M" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2171 | shows "infsum (\<lambda>x. Re (f x)) M = Re (infsum f M)" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2172 | by (simp add: assms has_sum_Re infsumI) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2173 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2174 | lemma summable_on_Re: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2175 | assumes "f summable_on M" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2176 | shows "(\<lambda>x. Re (f x)) summable_on M" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2177 | by (metis assms has_sum_Re summable_on_def) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2178 | |
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2179 | lemma has_sum_Im: | 
| 77359 | 2180 | assumes "(f has_sum a) M" | 
| 2181 | shows "((\<lambda>x. Im (f x)) has_sum Im a) M" | |
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2182 | using has_sum_comm_additive[where f=Im] | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2183 | using assms tendsto_Im by (fastforce simp add: o_def additive_def) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2184 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2185 | lemma infsum_Im: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2186 | assumes "f summable_on M" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2187 | shows "infsum (\<lambda>x. Im (f x)) M = Im (infsum f M)" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2188 | by (simp add: assms has_sum_Im infsumI) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2189 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2190 | lemma summable_on_Im: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2191 | assumes "f summable_on M" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2192 | shows "(\<lambda>x. Im (f x)) summable_on M" | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2193 | by (metis assms has_sum_Im summable_on_def) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2194 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2195 | lemma nonneg_infsum_le_0D_complex: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2196 | fixes f :: "'a \<Rightarrow> complex" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2197 | assumes "infsum f A \<le> 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2198 | and abs_sum: "f summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2199 | and nneg: "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2200 | and "x \<in> A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2201 | shows "f x = 0" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2202 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2203 | have \<open>Im (f x) = 0\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2204 | using assms(4) less_eq_complex_def nneg by auto | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2205 | moreover have \<open>Re (f x) = 0\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2206 | using assms by (auto simp add: summable_on_Re infsum_Re less_eq_complex_def intro: nonneg_infsum_le_0D[where A=A]) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2207 | ultimately show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2208 | by (simp add: complex_eqI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2209 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2210 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2211 | lemma nonneg_has_sum_le_0D_complex: | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2212 | fixes f :: "'a \<Rightarrow> complex" | 
| 77359 | 2213 | assumes "(f has_sum a) A" and \<open>a \<le> 0\<close> | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2214 | and "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> 0" and "x \<in> A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2215 | shows "f x = 0" | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2216 | by (metis assms infsumI nonneg_infsum_le_0D_complex summable_on_def) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2217 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2218 | text \<open>The lemma @{thm [source] infsum_mono_neutral} above applies to various linear ordered monoids such as the reals but not to the complex numbers.
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2219 | Thus we have a separate corollary for those:\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2220 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2221 | lemma infsum_mono_neutral_complex: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2222 | fixes f :: "'a \<Rightarrow> complex" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2223 | assumes [simp]: "f summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2224 | and [simp]: "g summable_on B" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2225 | assumes \<open>\<And>x. x \<in> A\<inter>B \<Longrightarrow> f x \<le> g x\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2226 | assumes \<open>\<And>x. x \<in> A-B \<Longrightarrow> f x \<le> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2227 | assumes \<open>\<And>x. x \<in> B-A \<Longrightarrow> g x \<ge> 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2228 | shows \<open>infsum f A \<le> infsum g B\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2229 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2230 | have \<open>infsum (\<lambda>x. Re (f x)) A \<le> infsum (\<lambda>x. Re (g x)) B\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2231 | by (smt (verit) assms infsum_cong infsum_mono_neutral less_eq_complex_def summable_on_Re zero_complex.simps(1)) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2232 | then have Re: \<open>Re (infsum f A) \<le> Re (infsum g B)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2233 | by (metis assms(1-2) infsum_Re) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2234 | have \<open>infsum (\<lambda>x. Im (f x)) A = infsum (\<lambda>x. Im (g x)) B\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2235 | by (smt (verit, best) assms(3-5) infsum_cong_neutral less_eq_complex_def zero_complex.simps(2)) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2236 | then have Im: \<open>Im (infsum f A) = Im (infsum g B)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2237 | by (metis assms(1-2) infsum_Im) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2238 | from Re Im show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2239 | by (auto simp: less_eq_complex_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2240 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2241 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2242 | lemma infsum_mono_complex: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2243 |   \<comment> \<open>For \<^typ>\<open>real\<close>, @{thm [source] infsum_mono} can be used. 
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2244 | But \<^typ>\<open>complex\<close> does not have the right typeclass.\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2245 | fixes f g :: "'a \<Rightarrow> complex" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2246 | assumes f_sum: "f summable_on A" and g_sum: "g summable_on A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2247 | assumes leq: "\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2248 | shows "infsum f A \<le> infsum g A" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2249 | by (metis DiffE IntD1 f_sum g_sum infsum_mono_neutral_complex leq) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2250 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2251 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2252 | lemma infsum_nonneg_complex: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2253 | fixes f :: "'a \<Rightarrow> complex" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2254 | assumes "f summable_on M" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2255 | and "\<And>x. x \<in> M \<Longrightarrow> 0 \<le> f x" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2256 | shows "infsum f M \<ge> 0" (is "?lhs \<ge> _") | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2257 | by (metis assms infsum_0_simp summable_on_0_simp infsum_mono_complex) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2258 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2259 | lemma infsum_cmod: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2260 | assumes "f summable_on M" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2261 | and fnn: "\<And>x. x \<in> M \<Longrightarrow> 0 \<le> f x" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2262 | shows "infsum (\<lambda>x. cmod (f x)) M = cmod (infsum f M)" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2263 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2264 | have \<open>complex_of_real (infsum (\<lambda>x. cmod (f x)) M) = infsum (\<lambda>x. complex_of_real (cmod (f x))) M\<close> | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2265 | proof (rule infsum_comm_additive[symmetric, unfolded o_def]) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2266 | have "(\<lambda>z. Re (f z)) summable_on M" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2267 | using assms summable_on_Re by blast | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2268 | also have "?this \<longleftrightarrow> f abs_summable_on M" | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2269 | using fnn by (intro summable_on_cong) (auto simp: less_eq_complex_def cmod_def) | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2270 | finally show \<dots> . | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2271 | qed (auto simp: additive_def) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2272 | also have \<open>\<dots> = infsum f M\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2273 | using fnn cmod_eq_Re complex_is_Real_iff less_eq_complex_def by (force cong: infsum_cong) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2274 | finally show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2275 | by (metis abs_of_nonneg infsum_def le_less_trans norm_ge_zero norm_infsum_bound norm_of_real not_le order_refl) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2276 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2277 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2278 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2279 | lemma summable_on_iff_abs_summable_on_complex: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2280 | fixes f :: \<open>'a \<Rightarrow> complex\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2281 | shows \<open>f summable_on A \<longleftrightarrow> f abs_summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2282 | proof (rule iffI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2283 | assume \<open>f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2284 | define i r ni nr n where \<open>i x = Im (f x)\<close> and \<open>r x = Re (f x)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2285 | and \<open>ni x = norm (i x)\<close> and \<open>nr x = norm (r x)\<close> and \<open>n x = norm (f x)\<close> for x | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2286 | from \<open>f summable_on A\<close> have \<open>i summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2287 | by (simp add: i_def[abs_def] summable_on_Im) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2288 | then have [simp]: \<open>ni summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2289 | using ni_def[abs_def] summable_on_iff_abs_summable_on_real by force | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2290 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2291 | from \<open>f summable_on A\<close> have \<open>r summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2292 | by (simp add: r_def[abs_def] summable_on_Re) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2293 | then have [simp]: \<open>nr summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2294 | by (metis nr_def summable_on_cong summable_on_iff_abs_summable_on_real) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2295 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2296 | have n_sum: \<open>n x \<le> nr x + ni x\<close> for x | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2297 | by (simp add: n_def nr_def ni_def r_def i_def cmod_le) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2298 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2299 | have *: \<open>(\<lambda>x. nr x + ni x) summable_on A\<close> | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2300 | by (simp add: summable_on_add) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2301 |   have "bdd_above (sum n ` {F. F \<subseteq> A \<and> finite F})"
 | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2302 | apply (rule bdd_aboveI[where M=\<open>infsum (\<lambda>x. nr x + ni x) A\<close>]) | 
| 76940 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2303 | using * n_sum by (auto simp flip: infsum_finite simp: ni_def nr_def intro!: infsum_mono_neutral) | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2304 | then show \<open>n summable_on A\<close> | 
| 
711cef61c0ce
Substantial de-applying and streamlining
 paulson <lp15@cam.ac.uk> parents: 
75462diff
changeset | 2305 | by (simp add: n_def nonneg_bdd_above_summable_on) | 
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2306 | next | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2307 | show \<open>f abs_summable_on A \<Longrightarrow> f summable_on A\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2308 | using abs_summable_summable by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2309 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2310 | |
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2311 | lemma summable_countable_complex: | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2312 | fixes f :: \<open>'a \<Rightarrow> complex\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2313 | assumes \<open>f summable_on A\<close> | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2314 |   shows \<open>countable {x\<in>A. f x \<noteq> 0}\<close>
 | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2315 | using abs_summable_countable assms summable_on_iff_abs_summable_on_complex by blast | 
| 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2316 | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2317 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2318 | (* TODO: figure all this out *) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2319 | inductive (in topological_space) convergent_filter :: "'a filter \<Rightarrow> bool" where | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2320 | "F \<le> nhds x \<Longrightarrow> convergent_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2321 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2322 | lemma (in topological_space) convergent_filter_iff: "convergent_filter F \<longleftrightarrow> (\<exists>x. F \<le> nhds x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2323 | by (auto simp: convergent_filter.simps) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2324 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2325 | lemma (in uniform_space) cauchy_filter_mono: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2326 | "cauchy_filter F \<Longrightarrow> F' \<le> F \<Longrightarrow> cauchy_filter F'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2327 | unfolding cauchy_filter_def by (meson dual_order.trans prod_filter_mono) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2328 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2329 | lemma (in uniform_space) convergent_filter_cauchy: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2330 | assumes "convergent_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2331 | shows "cauchy_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2332 | using assms cauchy_filter_mono nhds_imp_cauchy_filter[OF order.refl] | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2333 | by (auto simp: convergent_filter_iff) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2334 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2335 | lemma (in topological_space) convergent_filter_bot [simp, intro]: "convergent_filter bot" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2336 | by (simp add: convergent_filter_iff) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2337 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2338 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2339 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2340 | class complete_uniform_space = uniform_space + | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2341 | assumes cauchy_filter_convergent': "cauchy_filter (F :: 'a filter) \<Longrightarrow> F \<noteq> bot \<Longrightarrow> convergent_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2342 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2343 | lemma (in complete_uniform_space) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2344 | cauchy_filter_convergent: "cauchy_filter (F :: 'a filter) \<Longrightarrow> convergent_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2345 | using cauchy_filter_convergent'[of F] by (cases "F = bot") auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2346 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2347 | lemma (in complete_uniform_space) convergent_filter_iff_cauchy: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2348 | "convergent_filter F \<longleftrightarrow> cauchy_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2349 | using convergent_filter_cauchy cauchy_filter_convergent by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2350 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2351 | definition countably_generated_filter :: "'a filter \<Rightarrow> bool" where | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2352 | "countably_generated_filter F \<longleftrightarrow> (\<exists>U :: nat \<Rightarrow> 'a set. F = (INF (n::nat). principal (U n)))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2353 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2354 | lemma countably_generated_filter_has_antimono_basis: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2355 | assumes "countably_generated_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2356 | obtains B :: "nat \<Rightarrow> 'a set" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2357 | where "antimono B" and "F = (INF n. principal (B n))" and | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2358 | "\<And>P. eventually P F \<longleftrightarrow> (\<exists>i. \<forall>x\<in>B i. P x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2359 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2360 | from assms obtain B where B: "F = (INF (n::nat). principal (B n))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2361 | unfolding countably_generated_filter_def by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2362 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2363 | define B' where "B' = (\<lambda>n. \<Inter>k\<le>n. B k)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2364 | have "antimono B'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2365 | unfolding decseq_def B'_def by force | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2366 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2367 |   have "(INF n. principal (B' n)) = (INF n. INF k\<in>{..n}. principal (B k))"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2368 | unfolding B'_def by (intro INF_cong refl INF_principal_finite [symmetric]) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2369 | also have "\<dots> = (INF (n::nat). principal (B n))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2370 | apply (intro antisym) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2371 | apply (meson INF_lower INF_mono atMost_iff order_refl) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2372 | apply (meson INF_greatest INF_lower UNIV_I) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2373 | done | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2374 | also have "\<dots> = F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2375 | by (simp add: B) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2376 | finally have F: "F = (INF n. principal (B' n))" .. | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2377 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2378 | moreover have "eventually P F \<longleftrightarrow> (\<exists>i. eventually P (principal (B' i)))" for P | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2379 | unfolding F using \<open>antimono B'\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2380 | apply (subst eventually_INF_base) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2381 | apply (auto simp: decseq_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2382 | by (meson nat_le_linear) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2383 | ultimately show ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2384 | using \<open>antimono B'\<close> that[of B'] unfolding eventually_principal by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2385 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2386 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2387 | lemma (in uniform_space) cauchy_filter_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2388 | "cauchy_filter F \<longleftrightarrow> (\<forall>P. eventually P uniformity \<longrightarrow> (\<exists>X. eventually (\<lambda>x. x \<in> X) F \<and> (\<forall>z\<in>X\<times>X. P z)))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2389 | unfolding cauchy_filter_def le_filter_def | 
| 79529 
cb933e165dc3
tuned proof: avoid z3 to make it work on arm64-linux;
 wenzelm parents: 
77388diff
changeset | 2390 | apply (auto simp: eventually_prod_same) | 
| 
cb933e165dc3
tuned proof: avoid z3 to make it work on arm64-linux;
 wenzelm parents: 
77388diff
changeset | 2391 | apply (metis (full_types) eventually_mono mem_Collect_eq) | 
| 
cb933e165dc3
tuned proof: avoid z3 to make it work on arm64-linux;
 wenzelm parents: 
77388diff
changeset | 2392 | apply blast | 
| 
cb933e165dc3
tuned proof: avoid z3 to make it work on arm64-linux;
 wenzelm parents: 
77388diff
changeset | 2393 | done | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2394 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2395 | lemma (in uniform_space) controlled_sequences_convergent_imp_complete_aux_sequence: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2396 |   fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2397 | fixes F :: "'a filter" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2398 | assumes "cauchy_filter F" "F \<noteq> bot" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2399 | assumes "\<And>n. eventually (\<lambda>z. z \<in> U n) uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2400 | obtains g G where | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2401 | "antimono G" "\<And>n. g n \<in> G n" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2402 | "\<And>n. eventually (\<lambda>x. x \<in> G n) F" "\<And>n. G n \<times> G n \<subseteq> U n" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2403 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2404 | have "\<exists>C. eventually (\<lambda>x. x \<in> C) F \<and> C \<times> C \<subseteq> U n" for n | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2405 | using assms(1) assms(3)[of n] unfolding cauchy_filter_iff by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2406 | then obtain G where G: "\<And>n. eventually (\<lambda>x. x \<in> G n) F" "\<And>n. G n \<times> G n \<subseteq> U n" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2407 | by metis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2408 | define G' where "G' = (\<lambda>n. \<Inter>k\<le>n. G k)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2409 | have 1: "eventually (\<lambda>x. x \<in> G' n) F" for n | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2410 | using G by (auto simp: G'_def intro: eventually_ball_finite) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2411 | have 2: "G' n \<times> G' n \<subseteq> U n" for n | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2412 | using G unfolding G'_def by fast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2413 | have 3: "antimono G'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2414 | unfolding G'_def decseq_def by force | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2415 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2416 | have "\<exists>g. g \<in> G' n" for n | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2417 | using 1 assms(2) eventually_happens' by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2418 | then obtain g where g: "\<And>n. g n \<in> G' n" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2419 | by metis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2420 | from g 1 2 3 that[of G' g] show ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2421 | by metis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2422 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2423 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2424 | definition lift_filter :: "('a set \<Rightarrow> 'b filter) \<Rightarrow> 'a filter \<Rightarrow> 'b filter" where
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2425 |   "lift_filter f F = (INF X\<in>{X. eventually (\<lambda>x. x \<in> X) F}. f X)"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2426 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2427 | lemma lift_filter_top [simp]: "lift_filter g top = g UNIV" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2428 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2429 |   have "{X. \<forall>x::'b. x \<in> X} = {UNIV}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2430 | by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2431 | thus ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2432 | by (simp add: lift_filter_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2433 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2434 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2435 | lemma eventually_lift_filter_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2436 | assumes "mono g" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2437 | shows "eventually P (lift_filter g F) \<longleftrightarrow> (\<exists>X. eventually (\<lambda>x. x \<in> X) F \<and> eventually P (g X))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2438 | unfolding lift_filter_def | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2439 | proof (subst eventually_INF_base, goal_cases) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2440 | case 1 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2441 | thus ?case by (auto intro: exI[of _ UNIV]) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2442 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2443 | case (2 X Y) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2444 | thus ?case | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2445 | by (auto intro!: exI[of _ "X \<inter> Y"] eventually_conj monoD[OF assms]) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2446 | qed auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2447 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2448 | lemma lift_filter_le: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2449 | assumes "eventually (\<lambda>x. x \<in> X) F" "g X \<le> F'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2450 | shows "lift_filter g F \<le> F'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2451 | unfolding lift_filter_def | 
| 77388 | 2452 | by (metis INF_lower2 assms mem_Collect_eq) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2453 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2454 | definition lift_filter' :: "('a set \<Rightarrow> 'b set) \<Rightarrow> 'a filter \<Rightarrow> 'b filter" where
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2455 | "lift_filter' f F = lift_filter (principal \<circ> f) F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2456 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2457 | lemma lift_filter'_top [simp]: "lift_filter' g top = principal (g UNIV)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2458 | by (simp add: lift_filter'_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2459 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2460 | lemma eventually_lift_filter'_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2461 | assumes "mono g" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2462 | shows "eventually P (lift_filter' g F) \<longleftrightarrow> (\<exists>X. eventually (\<lambda>x. x \<in> X) F \<and> (\<forall>x\<in>g X. P x))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2463 | unfolding lift_filter'_def using assms | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2464 | by (subst eventually_lift_filter_iff) (auto simp: mono_def eventually_principal) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2465 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2466 | lemma lift_filter'_le: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2467 | assumes "eventually (\<lambda>x. x \<in> X) F" "principal (g X) \<le> F'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2468 | shows "lift_filter' g F \<le> F'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2469 | unfolding lift_filter'_def using assms | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2470 | by (intro lift_filter_le[where X = X]) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2471 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2472 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2473 | lemma (in uniform_space) comp_uniformity_le_uniformity: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2474 | "lift_filter' (\<lambda>X. X O X) uniformity \<le> uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2475 | unfolding le_filter_def | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2476 | proof safe | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2477 | fix P assume P: "eventually P uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2478 |   have [simp]: "mono (\<lambda>X::('a \<times> 'a) set. X O X)"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2479 | by (intro monoI) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2480 | from P obtain P' where P': "eventually P' uniformity " "(\<And>x y z. P' (x, y) \<Longrightarrow> P' (y, z) \<Longrightarrow> P (x, z))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2481 | using uniformity_transE by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2482 | show "eventually P (lift_filter' (\<lambda>X. X O X) uniformity)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2483 |     by (auto simp: eventually_lift_filter'_iff intro!: exI[of _ "{x. P' x}"] P')
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2484 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2485 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2486 | lemma (in uniform_space) comp_mem_uniformity_sets: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2487 | assumes "eventually (\<lambda>z. z \<in> X) uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2488 | obtains Y where "eventually (\<lambda>z. z \<in> Y) uniformity" "Y O Y \<subseteq> X" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2489 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2490 |   have [simp]: "mono (\<lambda>X::('a \<times> 'a) set. X O X)"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2491 | by (intro monoI) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2492 | have "eventually (\<lambda>z. z \<in> X) (lift_filter' (\<lambda>X. X O X) uniformity)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2493 | using assms comp_uniformity_le_uniformity using filter_leD by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2494 | thus ?thesis using that | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2495 | by (auto simp: eventually_lift_filter'_iff) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2496 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2497 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2498 | lemma (in uniform_space) le_nhds_of_cauchy_adhp_aux: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2499 | assumes "\<And>P. eventually P uniformity \<Longrightarrow> (\<exists>X. eventually (\<lambda>y. y \<in> X) F \<and> (\<forall>z\<in>X \<times> X. P z) \<and> (\<exists>y. P (x, y) \<and> y \<in> X))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2500 | shows "F \<le> nhds x" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2501 | unfolding le_filter_def | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2502 | proof safe | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2503 | fix P assume "eventually P (nhds x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2504 |   hence "\<forall>\<^sub>F z in uniformity. z \<in> {z. fst z = x \<longrightarrow> P (snd z)}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2505 | by (simp add: eventually_nhds_uniformity case_prod_unfold) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2506 |   then obtain Y where Y: "\<forall>\<^sub>F z in uniformity. z \<in> Y" "Y O Y \<subseteq> {z. fst z = x \<longrightarrow> P (snd z)}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2507 | using comp_mem_uniformity_sets by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2508 | obtain X y where Xy: "eventually (\<lambda>y. y \<in> X) F" "X\<times>X \<subseteq> Y" "(x, y) \<in> Y" "y \<in> X" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2509 | using assms[OF Y(1)] by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2510 | have *: "P x" if "x \<in> X" for x | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2511 | using Y(2) Xy(2-4) that unfolding relcomp_unfold by force | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2512 | show "eventually P F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2513 | by (rule eventually_mono[OF Xy(1)]) (use * in auto) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2514 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2515 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2516 | lemma (in uniform_space) eventually_uniformity_imp_nhds: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2517 | assumes "eventually P uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2518 | shows "eventually (\<lambda>y. P (x, y)) (nhds x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2519 | using assms unfolding eventually_nhds_uniformity by (elim eventually_mono) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2520 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2521 | lemma (in uniform_space) controlled_sequences_convergent_imp_complete_aux: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2522 |   fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2523 |   assumes gen: "countably_generated_filter (uniformity :: ('a \<times> 'a) filter)"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2524 | assumes U: "\<And>n. eventually (\<lambda>z. z \<in> U n) uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2525 | assumes conv: "\<And>(u :: nat \<Rightarrow> 'a). (\<And>N m n. N \<le> m \<Longrightarrow> N \<le> n \<Longrightarrow> (u m, u n) \<in> U N) \<Longrightarrow> convergent u" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2526 | assumes "cauchy_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2527 | shows "convergent_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2528 | proof (cases "F = bot") | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2529 | case False | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2530 | note F = \<open>cauchy_filter F\<close> \<open>F \<noteq> bot\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2531 |   from gen obtain B :: "nat \<Rightarrow> ('a \<times> 'a) set" where B:
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2532 | "antimono B" "uniformity = (INF n. principal (B n))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2533 | "\<And>P. eventually P uniformity \<longleftrightarrow> (\<exists>i. \<forall>x\<in>B i. P x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2534 | using countably_generated_filter_has_antimono_basis by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2535 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2536 | have ev_B: "eventually (\<lambda>z. z \<in> B n) uniformity" for n | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2537 | by (subst B(3)) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2538 | hence ev_B': "eventually (\<lambda>z. z \<in> B n \<inter> U n) uniformity" for n | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2539 | using U by (auto intro: eventually_conj) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2540 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2541 | obtain g G where gG: "antimono G" "\<And>n. g n \<in> G n" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2542 | "\<And>n. eventually (\<lambda>x. x \<in> G n) F" "\<And>n. G n \<times> G n \<subseteq> B n \<inter> U n" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2543 | using controlled_sequences_convergent_imp_complete_aux_sequence[of F "\<lambda>n. B n \<inter> U n", OF F ev_B'] | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2544 | by metis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2545 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2546 | have "convergent g" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2547 | proof (rule conv) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2548 | fix N m n :: nat | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2549 | assume mn: "N \<le> m" "N \<le> n" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2550 | have "(g m, g n) \<in> G m \<times> G n" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2551 | using gG by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2552 | also from mn have "\<dots> \<subseteq> G N \<times> G N" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2553 | by (intro Sigma_mono gG antimonoD[OF gG(1)]) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2554 | also have "\<dots> \<subseteq> U N" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2555 | using gG by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2556 | finally show "(g m, g n) \<in> U N" . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2557 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2558 | then obtain L where G: "g \<longlonglongrightarrow> L" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2559 | unfolding convergent_def by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2560 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2561 | have "F \<le> nhds L" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2562 | proof (rule le_nhds_of_cauchy_adhp_aux) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2563 | fix P :: "'a \<times> 'a \<Rightarrow> bool" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2564 | assume P: "eventually P uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2565 | hence "eventually (\<lambda>n. \<forall>x\<in>B n. P x) sequentially" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2566 | using \<open>antimono B\<close> unfolding B(3) eventually_sequentially decseq_def by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2567 | moreover have "eventually (\<lambda>n. P (L, g n)) sequentially" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2568 | using P eventually_compose_filterlim eventually_uniformity_imp_nhds G by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2569 | ultimately have "eventually (\<lambda>n. (\<forall>x\<in>B n. P x) \<and> P (L, g n)) sequentially" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2570 | by eventually_elim auto | 
| 77388 | 2571 | then obtain n where "\<forall>x\<in>B n. P x" "P (L, g n)" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2572 | unfolding eventually_at_top_linorder by blast | 
| 77388 | 2573 | then show "\<exists>X. (\<forall>\<^sub>F y in F. y \<in> X) \<and> (\<forall>z\<in>X \<times> X. P z) \<and> (\<exists>y. P (L, y) \<and> y \<in> X)" | 
| 2574 | using gG by blast+ | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2575 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2576 | thus "convergent_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2577 | by (auto simp: convergent_filter_iff) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2578 | qed auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2579 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2580 | theorem (in uniform_space) controlled_sequences_convergent_imp_complete: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2581 |   fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2582 |   assumes gen: "countably_generated_filter (uniformity :: ('a \<times> 'a) filter)"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2583 | assumes U: "\<And>n. eventually (\<lambda>z. z \<in> U n) uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2584 | assumes conv: "\<And>(u :: nat \<Rightarrow> 'a). (\<And>N m n. N \<le> m \<Longrightarrow> N \<le> n \<Longrightarrow> (u m, u n) \<in> U N) \<Longrightarrow> convergent u" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2585 | shows "class.complete_uniform_space open uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2586 | by unfold_locales (use assms controlled_sequences_convergent_imp_complete_aux in blast) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2587 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2588 | lemma filtermap_prod_filter: "filtermap (map_prod f g) (F \<times>\<^sub>F G) = filtermap f F \<times>\<^sub>F filtermap g G" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2589 | proof (intro antisym) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2590 | show "filtermap (map_prod f g) (F \<times>\<^sub>F G) \<le> filtermap f F \<times>\<^sub>F filtermap g G" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2591 | by (auto simp: le_filter_def eventually_filtermap eventually_prod_filter) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2592 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2593 | show "filtermap f F \<times>\<^sub>F filtermap g G \<le> filtermap (map_prod f g) (F \<times>\<^sub>F G)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2594 | unfolding le_filter_def | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2595 | proof safe | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2596 | fix P assume P: "eventually P (filtermap (map_prod f g) (F \<times>\<^sub>F G))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2597 | then obtain Pf Pg where *: "eventually Pf F" "eventually Pg G" "\<forall>x. Pf x \<longrightarrow> (\<forall>y. Pg y \<longrightarrow> P (f x, g y))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2598 | by (auto simp: eventually_filtermap eventually_prod_filter) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2599 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2600 | define Pf' where "Pf' = (\<lambda>x. \<exists>y. x = f y \<and> Pf y)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2601 | define Pg' where "Pg' = (\<lambda>x. \<exists>y. x = g y \<and> Pg y)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2602 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2603 | from *(1) have "\<forall>\<^sub>F x in F. Pf' (f x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2604 | by eventually_elim (auto simp: Pf'_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2605 | moreover from *(2) have "\<forall>\<^sub>F x in G. Pg' (g x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2606 | by eventually_elim (auto simp: Pg'_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2607 | moreover have "(\<forall>x y. Pf' x \<longrightarrow> Pg' y \<longrightarrow> P (x, y))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2608 | using *(3) by (auto simp: Pf'_def Pg'_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2609 | ultimately show "eventually P (filtermap f F \<times>\<^sub>F filtermap g G)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2610 | unfolding eventually_prod_filter eventually_filtermap | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2611 | by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2612 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2613 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2614 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2615 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2616 | lemma (in uniform_space) Cauchy_seq_iff_tendsto: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2617 | "Cauchy f \<longleftrightarrow> filterlim (map_prod f f) uniformity (at_top \<times>\<^sub>F at_top)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2618 | unfolding Cauchy_uniform cauchy_filter_def filterlim_def filtermap_prod_filter .. | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2619 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2620 | theorem (in uniform_space) controlled_seq_imp_Cauchy_seq: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2621 |   fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2622 | assumes U: "\<And>P. eventually P uniformity \<Longrightarrow> (\<exists>n. \<forall>x\<in>U n. P x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2623 | assumes controlled: "\<And>N m n. N \<le> m \<Longrightarrow> N \<le> n \<Longrightarrow> (f m, f n) \<in> U N" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2624 | shows "Cauchy f" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2625 | unfolding Cauchy_seq_iff_tendsto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2626 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2627 | show "filterlim (map_prod f f) uniformity (sequentially \<times>\<^sub>F sequentially)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2628 | unfolding filterlim_def le_filter_def | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2629 | proof safe | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2630 | fix P :: "'a \<times> 'a \<Rightarrow> bool" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2631 | assume P: "eventually P uniformity" | 
| 77388 | 2632 | from U[OF this] obtain N where "\<forall>x\<in>U N. P x" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2633 | by blast | 
| 77388 | 2634 | then show "eventually P (filtermap (map_prod f f) (sequentially \<times>\<^sub>F sequentially))" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2635 | unfolding eventually_filtermap eventually_prod_sequentially | 
| 77388 | 2636 | by (metis controlled map_prod_simp) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2637 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2638 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2639 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2640 | lemma (in uniform_space) Cauchy_seq_convergent_imp_complete_aux: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2641 |   fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2642 |   assumes gen: "countably_generated_filter (uniformity :: ('a \<times> 'a) filter)"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2643 | assumes conv: "\<And>(u :: nat \<Rightarrow> 'a). Cauchy u \<Longrightarrow> convergent u" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2644 | assumes "cauchy_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2645 | shows "convergent_filter F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2646 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2647 |   from gen obtain B :: "nat \<Rightarrow> ('a \<times> 'a) set" where B:
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2648 | "antimono B" "uniformity = (INF n. principal (B n))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2649 | "\<And>P. eventually P uniformity \<longleftrightarrow> (\<exists>i. \<forall>x\<in>B i. P x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2650 | using countably_generated_filter_has_antimono_basis by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2651 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2652 | show ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2653 | proof (rule controlled_sequences_convergent_imp_complete_aux[where U = B]) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2654 | show "\<forall>\<^sub>F z in uniformity. z \<in> B n" for n | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2655 | unfolding B(3) by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2656 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2657 | fix f :: "nat \<Rightarrow> 'a" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2658 | assume f: "\<And>N m n. N \<le> m \<Longrightarrow> N \<le> n \<Longrightarrow> (f m, f n) \<in> B N" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2659 | have "Cauchy f" using f B | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2660 | by (intro controlled_seq_imp_Cauchy_seq[where U = B]) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2661 | with conv show "convergent f" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2662 | by simp | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2663 | qed fact+ | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2664 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2665 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2666 | theorem (in uniform_space) Cauchy_seq_convergent_imp_complete: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2667 |   fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2668 |   assumes gen: "countably_generated_filter (uniformity :: ('a \<times> 'a) filter)"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2669 | assumes conv: "\<And>(u :: nat \<Rightarrow> 'a). Cauchy u \<Longrightarrow> convergent u" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2670 | shows "class.complete_uniform_space open uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2671 | by unfold_locales (use assms Cauchy_seq_convergent_imp_complete_aux in blast) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2672 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2673 | lemma (in metric_space) countably_generated_uniformity: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2674 | "countably_generated_filter uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2675 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2676 |   have "(INF e\<in>{0<..}. principal {(x, y). dist (x::'a) y < e}) =
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2677 |         (INF n\<in>UNIV. principal {(x, y). dist x y < 1 / real (Suc n)})" (is "?F = ?G")
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2678 | unfolding uniformity_dist | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2679 | proof (intro antisym) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2680 |     have "?G = (INF e\<in>(\<lambda>n. 1 / real (Suc n)) ` UNIV. principal {(x, y). dist x y < e})"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2681 | by (simp add: image_image) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2682 | also have "\<dots> \<ge> ?F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2683 | by (intro INF_superset_mono) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2684 | finally show "?F \<le> ?G" . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2685 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2686 | show "?G \<le> ?F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2687 | unfolding le_filter_def | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2688 | proof safe | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2689 | fix P assume "eventually P ?F" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2690 |       then obtain \<epsilon> where \<epsilon>: "\<epsilon> > 0" "eventually P (principal {(x, y). dist x y < \<epsilon>})"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2691 | proof (subst (asm) eventually_INF_base, goal_cases) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2692 | case (2 \<epsilon>1 \<epsilon>2) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2693 | thus ?case | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2694 | by (intro bexI[of _ "min \<epsilon>1 \<epsilon>2"]) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2695 | qed auto | 
| 77388 | 2696 | from \<open>\<epsilon> > 0\<close> obtain n where "1 / real (Suc n) < \<epsilon>" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2697 | using nat_approx_posE by blast | 
| 77388 | 2698 |       then have "eventually P (principal {(x, y). dist x y < 1 / real (Suc n)})"
 | 
| 2699 | using \<epsilon>(2) by (auto simp: eventually_principal) | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2700 | thus "eventually P ?G" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2701 | by (intro eventually_INF1) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2702 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2703 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2704 | thus "countably_generated_filter uniformity" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2705 | unfolding countably_generated_filter_def uniformity_dist by fast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2706 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2707 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2708 | subclass (in complete_space) complete_uniform_space | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2709 | proof (rule Cauchy_seq_convergent_imp_complete) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2710 | show "convergent f" if "Cauchy f" for f | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2711 | using Cauchy_convergent that by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2712 | qed (fact countably_generated_uniformity) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2713 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2714 | lemma (in complete_uniform_space) complete_UNIV_cuspace [intro]: "complete UNIV" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2715 | unfolding complete_uniform using cauchy_filter_convergent | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2716 | by (auto simp: convergent_filter.simps) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2717 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2718 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2719 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2720 | lemma norm_infsum_le: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2721 | assumes "(f has_sum S) X" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2722 | assumes "(g has_sum T) X" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2723 | assumes "\<And>x. x \<in> X \<Longrightarrow> norm (f x) \<le> g x" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2724 | shows "norm S \<le> T" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2725 | proof (rule tendsto_le) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2726 | show "((\<lambda>Y. norm (\<Sum>x\<in>Y. f x)) \<longlongrightarrow> norm S) (finite_subsets_at_top X)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2727 | using assms(1) unfolding has_sum_def by (intro tendsto_norm) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2728 | show "((\<lambda>Y. \<Sum>x\<in>Y. g x) \<longlongrightarrow> T) (finite_subsets_at_top X)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2729 | using assms(2) unfolding has_sum_def . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2730 | show "\<forall>\<^sub>F x in finite_subsets_at_top X. norm (sum f x) \<le> (\<Sum>x\<in>x. g x)" | 
| 77388 | 2731 | by (simp add: assms(3) eventually_finite_subsets_at_top_weakI subsetD sum_norm_le) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2732 | qed auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2733 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2734 | (* | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2735 | lemma summable_on_Sigma: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2736 | fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2737 |     and f :: \<open>'a \<Rightarrow> 'b \<Rightarrow> 'c::{comm_monoid_add, t2_space, uniform_space}\<close>
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2738 | assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'c,y). x+y)\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2739 | assumes summableAB: "(\<lambda>(x,y). f x y) summable_on (Sigma A B)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2740 | assumes summableB: \<open>\<And>x. x\<in>A \<Longrightarrow> (f x) summable_on (B x)\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2741 | shows \<open>(\<lambda>x. infsum (f x) (B x)) summable_on A\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2742 | *) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2743 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2744 | lemma has_sum_imp_summable: "(f has_sum S) A \<Longrightarrow> f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2745 | by (auto simp: summable_on_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2746 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2747 | lemma has_sum_reindex_bij_betw: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2748 | assumes "bij_betw g A B" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2749 | shows "((\<lambda>x. f (g x)) has_sum S) A = (f has_sum S) B" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2750 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2751 | have "((\<lambda>x. f (g x)) has_sum S) A \<longleftrightarrow> (f has_sum S) (g ` A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2752 | by (subst has_sum_reindex) (use assms in \<open>auto dest: bij_betw_imp_inj_on simp: o_def\<close>) | 
| 77388 | 2753 | then show ?thesis | 
| 2754 | using assms bij_betw_imp_surj_on by blast | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2755 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2756 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2757 | lemma has_sum_reindex_bij_witness: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2758 | assumes "\<And>a. a \<in> S \<Longrightarrow> i (j a) = a" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2759 | assumes "\<And>a. a \<in> S \<Longrightarrow> j a \<in> T" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2760 | assumes "\<And>b. b \<in> T \<Longrightarrow> j (i b) = b" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2761 | assumes "\<And>b. b \<in> T \<Longrightarrow> i b \<in> S" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2762 | assumes "\<And>a. a \<in> S \<Longrightarrow> h (j a) = g a" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2763 | assumes "s = s'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2764 | shows "(g has_sum s) S = (h has_sum s') T" | 
| 77388 | 2765 | by (smt (verit, del_insts) assms bij_betwI' has_sum_cong has_sum_reindex_bij_betw) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2766 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2767 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2768 | lemma has_sum_homomorphism: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2769 | assumes "(f has_sum S) A" "h 0 = 0" "\<And>a b. h (a + b) = h a + h b" "continuous_on UNIV h" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2770 | shows "((\<lambda>x. h (f x)) has_sum (h S)) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2771 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2772 | have "sum (h \<circ> f) X = h (sum f X)" for X | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2773 | by (induction X rule: infinite_finite_induct) (simp_all add: assms) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2774 | hence sum_h: "sum (h \<circ> f) = h \<circ> sum f" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2775 | by (intro ext) auto | 
| 77388 | 2776 | then have "((h \<circ> f) has_sum h S) A" | 
| 2777 | using assms | |
| 2778 | by (metis UNIV_I continuous_on_def has_sum_comm_additive_general o_apply) | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2779 | thus ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2780 | by (simp add: o_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2781 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2782 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2783 | lemma summable_on_homomorphism: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2784 | assumes "f summable_on A" "h 0 = 0" "\<And>a b. h (a + b) = h a + h b" "continuous_on UNIV h" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2785 | shows "(\<lambda>x. h (f x)) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2786 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2787 | from assms(1) obtain S where "(f has_sum S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2788 | by (auto simp: summable_on_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2789 | hence "((\<lambda>x. h (f x)) has_sum h S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2790 | by (rule has_sum_homomorphism) (use assms in auto) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2791 | thus ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2792 | by (auto simp: summable_on_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2793 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2794 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2795 | lemma infsum_homomorphism_strong: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2796 |   fixes h :: "'a :: {t2_space, topological_comm_monoid_add} \<Rightarrow>
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2797 |                 'b :: {t2_space, topological_comm_monoid_add}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2798 | assumes "(\<lambda>x. h (f x)) summable_on A \<longleftrightarrow> f summable_on A" | 
| 77388 | 2799 | assumes "h 0 = 0" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2800 | assumes "\<And>S. (f has_sum S) A \<Longrightarrow> ((\<lambda>x. h (f x)) has_sum (h S)) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2801 | shows "infsum (\<lambda>x. h (f x)) A = h (infsum f A)" | 
| 77388 | 2802 | by (metis assms has_sum_infsum infsumI infsum_not_exists) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2803 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2804 | lemma has_sum_bounded_linear: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2805 | assumes "bounded_linear h" and "(f has_sum S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2806 | shows "((\<lambda>x. h (f x)) has_sum h S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2807 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2808 | interpret bounded_linear h by fact | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2809 | from assms(2) show ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2810 | by (rule has_sum_homomorphism) (auto simp: add intro!: continuous_on) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2811 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2812 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2813 | lemma summable_on_bounded_linear: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2814 | assumes "bounded_linear h" and "f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2815 | shows "(\<lambda>x. h (f x)) summable_on A" | 
| 77388 | 2816 | by (metis assms has_sum_bounded_linear summable_on_def) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2817 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2818 | lemma summable_on_bounded_linear_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2819 | assumes "bounded_linear h" and "bounded_linear h'" and "\<And>x. h' (h x) = x" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2820 | shows "(\<lambda>x. h (f x)) summable_on A \<longleftrightarrow> f summable_on A" | 
| 77388 | 2821 | by (metis (full_types) assms summable_on_bounded_linear summable_on_cong) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2822 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2823 | lemma infsum_bounded_linear_strong: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2824 | fixes h :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_vector" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2825 | assumes "(\<lambda>x. h (f x)) summable_on A \<longleftrightarrow> f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2826 | assumes "bounded_linear h" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2827 | shows "infsum (\<lambda>x. h (f x)) A = h (infsum f A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2828 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2829 | interpret bounded_linear h by fact | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2830 | show ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2831 | by (rule infsum_homomorphism_strong) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2832 | (insert assms, auto intro: add continuous_on has_sum_bounded_linear) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2833 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2834 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2835 | lemma infsum_bounded_linear_strong': | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2836 | fixes mult :: "'c :: zero \<Rightarrow> 'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_vector" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2837 | assumes "c \<noteq> 0 \<Longrightarrow> (\<lambda>x. mult c (f x)) summable_on A \<longleftrightarrow> f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2838 | assumes "bounded_linear (mult c)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2839 | assumes [simp]: "\<And>x. mult 0 x = 0" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2840 | shows "infsum (\<lambda>x. mult c (f x)) A = mult c (infsum f A)" | 
| 77388 | 2841 | by (metis assms infsum_0 infsum_bounded_linear_strong) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2842 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2843 | lemma has_sum_of_nat: "(f has_sum S) A \<Longrightarrow> ((\<lambda>x. of_nat (f x)) has_sum of_nat S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2844 | by (erule has_sum_homomorphism) (auto intro!: continuous_intros) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2845 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2846 | lemma has_sum_of_int: "(f has_sum S) A \<Longrightarrow> ((\<lambda>x. of_int (f x)) has_sum of_int S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2847 | by (erule has_sum_homomorphism) (auto intro!: continuous_intros) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2848 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2849 | lemma summable_on_of_nat: "f summable_on A \<Longrightarrow> (\<lambda>x. of_nat (f x)) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2850 | by (erule summable_on_homomorphism) (auto intro!: continuous_intros) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2851 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2852 | lemma summable_on_of_int: "f summable_on A \<Longrightarrow> (\<lambda>x. of_int (f x)) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2853 | by (erule summable_on_homomorphism) (auto intro!: continuous_intros) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2854 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2855 | lemma summable_on_discrete_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2856 |   fixes f :: "'a \<Rightarrow> 'b :: {discrete_topology, topological_comm_monoid_add, cancel_comm_monoid_add}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2857 |   shows "f summable_on A \<longleftrightarrow> finite {x\<in>A. f x \<noteq> 0}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2858 | proof | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2859 |   assume *: "finite {x\<in>A. f x \<noteq> 0}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2860 |   hence "f summable_on {x\<in>A. f x \<noteq> 0}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2861 | by (rule summable_on_finite) | 
| 77388 | 2862 | then show "f summable_on A" | 
| 2863 | by (smt (verit) DiffE mem_Collect_eq summable_on_cong_neutral) | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2864 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2865 | assume "f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2866 | then obtain S where "(f has_sum S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2867 | by (auto simp: summable_on_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2868 | hence "\<forall>\<^sub>F x in finite_subsets_at_top A. sum f x = S" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2869 | unfolding has_sum_def tendsto_discrete . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2870 | then obtain X where X: "finite X" "X \<subseteq> A" "\<And>Y. finite Y \<Longrightarrow> X \<subseteq> Y \<Longrightarrow> Y \<subseteq> A \<Longrightarrow> sum f Y = S" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2871 | unfolding eventually_finite_subsets_at_top by metis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2872 |   have "{x\<in>A. f x \<noteq> 0} \<subseteq> X"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2873 | proof | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2874 |     fix x assume x: "x \<in> {x\<in>A. f x \<noteq> 0}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2875 | show "x \<in> X" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2876 | proof (rule ccontr) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2877 | assume [simp]: "x \<notin> X" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2878 | have "sum f (insert x X) = S" | 
| 77388 | 2879 | using X x by (intro X) auto | 
| 2880 | then have "f x = 0" | |
| 2881 | using X by auto | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2882 | with x show False | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2883 | by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2884 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2885 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2886 |   thus "finite {x\<in>A. f x \<noteq> 0}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2887 | using X(1) finite_subset by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2888 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2889 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2890 | lemma has_sum_imp_sums: "(f has_sum S) (UNIV :: nat set) \<Longrightarrow> f sums S" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2891 | unfolding sums_def has_sum_def by (rule filterlim_compose[OF _ filterlim_lessThan_at_top]) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2892 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2893 | lemma summable_on_imp_summable: "f summable_on (UNIV :: nat set) \<Longrightarrow> summable f" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2894 | unfolding summable_on_def summable_def by (auto dest: has_sum_imp_sums) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2895 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2896 | lemma summable_on_UNIV_nonneg_real_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2897 | assumes "\<And>n. f n \<ge> (0 :: real)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2898 | shows "f summable_on UNIV \<longleftrightarrow> summable f" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2899 | using assms by (auto intro: norm_summable_imp_summable_on summable_on_imp_summable) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2900 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2901 | lemma summable_on_imp_bounded_partial_sums: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2902 |   fixes f :: "_ \<Rightarrow> 'a :: {topological_comm_monoid_add, linorder_topology}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2903 | assumes f: "f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2904 | shows "\<exists>C. eventually (\<lambda>X. sum f X \<le> C) (finite_subsets_at_top A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2905 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2906 | from assms obtain S where S: "(sum f \<longlongrightarrow> S) (finite_subsets_at_top A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2907 | unfolding summable_on_def has_sum_def by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2908 | show ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2909 | proof (cases "\<exists>C. C > S") | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2910 | case True | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2911 | then obtain C where C: "C > S" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2912 | by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2913 | have "\<forall>\<^sub>F X in finite_subsets_at_top A. sum f X < C" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2914 | using S C by (rule order_tendstoD(2)) | 
| 77388 | 2915 | thus ?thesis | 
| 2916 | by (meson eventually_mono nless_le) | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2917 | next | 
| 77388 | 2918 | case False thus ?thesis | 
| 2919 | by (meson not_eventuallyD not_le_imp_less) | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2920 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2921 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2922 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2923 | lemma has_sum_mono': | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2924 |   fixes S S' :: "'a :: {linorder_topology, ordered_comm_monoid_add, topological_comm_monoid_add}"
 | 
| 77388 | 2925 | assumes f: "(f has_sum S) A" "(f has_sum S') B" | 
| 2926 | and AB: "A \<subseteq> B" "\<And>x. x \<in> B - A \<Longrightarrow> f x \<ge> 0" | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2927 | shows "S \<le> S'" | 
| 77388 | 2928 | using AB has_sum_mono_neutral[OF f] by fastforce | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2929 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2930 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2931 | context | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2932 |   assumes "SORT_CONSTRAINT('a :: {topological_comm_monoid_add, order_topology,
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2933 | ordered_comm_monoid_add, conditionally_complete_linorder})" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2934 | begin | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2935 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2936 | text \<open> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2937 | Any family of non-negative numbers with bounded partial sums is summable, and the sum | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2938 | is simply the supremum of the partial sums. | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2939 | \<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2940 | lemma nonneg_bounded_partial_sums_imp_has_sum_SUP: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2941 | assumes nonneg: "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> (0::'a)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2942 | and bound: "eventually (\<lambda>X. sum f X \<le> C) (finite_subsets_at_top A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2943 |   shows   "(f has_sum (SUP X\<in>{X. X \<subseteq> A \<and> finite X}. sum f X)) A"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2944 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2945 | from bound obtain X0 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2946 | where X0: "X0 \<subseteq> A" "finite X0" "\<And>X. X0 \<subseteq> X \<Longrightarrow> X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> sum f X \<le> C" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2947 | by (force simp: eventually_finite_subsets_at_top) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2948 | have bound': "sum f X \<le> C" if "X \<subseteq> A" "finite X" for X | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2949 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2950 | have "sum f X \<le> sum f (X \<union> X0)" | 
| 77388 | 2951 | using that X0 assms(1) by (intro sum_mono2) auto | 
| 2952 | also have "\<dots> \<le> C" | |
| 2953 | by (simp add: X0 that) | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2954 | finally show ?thesis . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2955 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2956 |   hence bdd: "bdd_above (sum f ` {X. X \<subseteq> A \<and> finite X})"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2957 | by (auto simp: bdd_above_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2958 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2959 | show ?thesis unfolding has_sum_def | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2960 | proof (rule increasing_tendsto) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2961 |     show "\<forall>\<^sub>F X in finite_subsets_at_top A. sum f X \<le> Sup (sum f ` {X. X \<subseteq> A \<and> finite X})"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2962 | by (intro eventually_finite_subsets_at_top_weakI cSUP_upper[OF _ bdd]) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2963 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2964 |     fix y assume "y < Sup (sum f ` {X. X \<subseteq> A \<and> finite X})"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2965 | then obtain X where X: "X \<subseteq> A" "finite X" "y < sum f X" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2966 | by (subst (asm) less_cSUP_iff[OF _ bdd]) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2967 | from X have "eventually (\<lambda>X'. X \<subseteq> X' \<and> X' \<subseteq> A \<and> finite X') (finite_subsets_at_top A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2968 | by (auto simp: eventually_finite_subsets_at_top) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2969 | thus "eventually (\<lambda>X'. y < sum f X') (finite_subsets_at_top A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2970 | proof eventually_elim | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2971 | case (elim X') | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2972 | note \<open>y < sum f X\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2973 | also have "sum f X \<le> sum f X'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2974 | using nonneg elim by (intro sum_mono2) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2975 | finally show ?case . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2976 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2977 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2978 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2979 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2980 | lemma nonneg_bounded_partial_sums_imp_summable_on: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2981 | assumes nonneg: "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> (0::'a)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2982 | and bound: "eventually (\<lambda>X. sum f X \<le> C) (finite_subsets_at_top A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2983 | shows "f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2984 | using nonneg_bounded_partial_sums_imp_has_sum_SUP[OF assms] by (auto simp: summable_on_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2985 | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2986 | end | 
| 74791 
227915e07891
more material for HOL-Analysis.Infinite_Sum
 Manuel Eberl <manuel@pruvisto.org> parents: 
74642diff
changeset | 2987 | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2988 | context | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2989 |   assumes "SORT_CONSTRAINT('a :: {topological_comm_monoid_add, linorder_topology,
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2990 | ordered_comm_monoid_add, conditionally_complete_linorder})" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2991 | begin | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2992 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2993 | lemma summable_on_comparison_test: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2994 | assumes "f summable_on A" and "\<And>x. x \<in> A \<Longrightarrow> g x \<le> f x" and "\<And>x. x \<in> A \<Longrightarrow> (0::'a) \<le> g x" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2995 | shows "g summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2996 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2997 | obtain C where C: "\<forall>\<^sub>F X in finite_subsets_at_top A. sum f X \<le> C" | 
| 77388 | 2998 | using assms(1) summable_on_imp_bounded_partial_sums by blast | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 2999 | show ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3000 | proof (rule nonneg_bounded_partial_sums_imp_summable_on) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3001 | show "\<forall>\<^sub>F X in finite_subsets_at_top A. sum g X \<le> C" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3002 | using C assms | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3003 | unfolding eventually_finite_subsets_at_top | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3004 | by (smt (verit, ccfv_SIG) order_trans subsetD sum_mono) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3005 | qed (use assms in auto) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3006 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3007 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3008 | end | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3009 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3010 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3011 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3012 | lemma summable_on_subset: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3013 |   fixes f :: "_ \<Rightarrow> 'a :: {uniform_topological_group_add, topological_comm_monoid_add, ab_group_add, complete_uniform_space}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3014 | assumes "f summable_on A" "B \<subseteq> A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3015 | shows "f summable_on B" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3016 | by (rule summable_on_subset_aux[OF _ _ assms]) (auto simp: uniformly_continuous_add) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3017 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3018 | lemma summable_on_union: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3019 |   fixes f :: "_ \<Rightarrow> 'a :: {uniform_topological_group_add, topological_comm_monoid_add, ab_group_add, complete_uniform_space}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3020 | assumes "f summable_on A" "f summable_on B" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3021 | shows "f summable_on (A \<union> B)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3022 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3023 | have "f summable_on (A \<union> (B - A))" | 
| 77388 | 3024 | by (meson Diff_disjoint Diff_subset assms summable_on_Un_disjoint summable_on_subset) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3025 | also have "A \<union> (B - A) = A \<union> B" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3026 | by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3027 | finally show ?thesis . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3028 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3029 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3030 | lemma summable_on_insert_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3031 |   fixes f :: "_ \<Rightarrow> 'a :: {uniform_topological_group_add, topological_comm_monoid_add, ab_group_add, complete_uniform_space}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3032 | shows "f summable_on insert x A \<longleftrightarrow> f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3033 |   using summable_on_union[of f A "{x}"] by (auto intro: summable_on_subset)
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3034 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3035 | lemma has_sum_finiteI: "finite A \<Longrightarrow> S = sum f A \<Longrightarrow> (f has_sum S) A" | 
| 77388 | 3036 | by simp | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3037 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3038 | lemma has_sum_insert: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3039 | fixes f :: "'a \<Rightarrow> 'b :: topological_comm_monoid_add" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3040 | assumes "x \<notin> A" and "(f has_sum S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3041 | shows "(f has_sum (f x + S)) (insert x A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3042 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3043 |   have "(f has_sum (f x + S)) ({x} \<union> A)"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3044 | using assms by (intro has_sum_Un_disjoint) (auto intro: has_sum_finiteI) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3045 | thus ?thesis by simp | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3046 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3047 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3048 | lemma infsum_insert: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3049 |   fixes f :: "_ \<Rightarrow> 'a :: {topological_comm_monoid_add, t2_space}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3050 | assumes "f summable_on A" "a \<notin> A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3051 | shows "infsum f (insert a A) = f a + infsum f A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3052 | by (meson assms has_sum_insert infsumI summable_iff_has_sum_infsum) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3053 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3054 | lemma has_sum_SigmaD: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3055 |   fixes f :: "'b \<times> 'c \<Rightarrow> 'a :: {topological_comm_monoid_add,t3_space}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3056 | assumes sum1: "(f has_sum S) (Sigma A B)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3057 | assumes sum2: "\<And>x. x \<in> A \<Longrightarrow> ((\<lambda>y. f (x, y)) has_sum g x) (B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3058 | shows "(g has_sum S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3059 | unfolding has_sum_def tendsto_def eventually_finite_subsets_at_top | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3060 | proof (safe, goal_cases) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3061 | case (1 X) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3062 | with nhds_closed[of S X] obtain X' | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3063 | where X': "S \<in> X'" "closed X'" "X' \<subseteq> X" "eventually (\<lambda>y. y \<in> X') (nhds S)" by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3064 | from X'(4) obtain X'' where X'': "S \<in> X''" "open X''" "X'' \<subseteq> X'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3065 | by (auto simp: eventually_nhds) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3066 |   with sum1 obtain Y :: "('b \<times> 'c) set"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3067 | where Y: "Y \<subseteq> Sigma A B" "finite Y" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3068 | "\<And>Z. Y \<subseteq> Z \<Longrightarrow> Z \<subseteq> Sigma A B \<Longrightarrow> finite Z \<Longrightarrow> sum f Z \<in> X''" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3069 | unfolding has_sum_def tendsto_def eventually_finite_subsets_at_top by force | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3070 | define Y1 :: "'b set" where "Y1 = fst ` Y" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3071 | from Y have Y1: "Y1 \<subseteq> A" by (auto simp: Y1_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3072 |   define Y2 :: "'b \<Rightarrow> 'c set" where "Y2 = (\<lambda>x. {y. (x, y) \<in> Y})"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3073 | have Y2: "finite (Y2 x)" "Y2 x \<subseteq> B x" if "x \<in> A" for x | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3074 | using that Y(1,2) unfolding Y2_def | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3075 | by (force simp: image_iff intro: finite_subset[of _ "snd ` Y"])+ | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3076 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3077 | show ?case | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3078 | proof (rule exI[of _ Y1], safe, goal_cases) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3079 | case (3 Z) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3080 | define H where "H = (INF x\<in>Z. filtercomap (\<lambda>p. p x) (finite_subsets_at_top (B x)))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3081 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3082 | have "sum g Z \<in> X'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3083 | proof (rule Lim_in_closed_set) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3084 | show "closed X'" by fact | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3085 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3086 | show "((\<lambda>B'. sum (\<lambda>x. sum (\<lambda>y. f (x, y)) (B' x)) Z) \<longlongrightarrow> sum g Z) H" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3087 | unfolding H_def | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3088 | proof (intro tendsto_sum filterlim_INF') | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3089 | fix x assume x: "x \<in> Z" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3090 | with 3 have "x \<in> A" by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3091 | from sum2[OF this] have "(sum (\<lambda>y. f (x, y)) \<longlongrightarrow> g x) (finite_subsets_at_top (B x))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3092 | by (simp add: has_sum_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3093 | thus "((\<lambda>B'. sum (\<lambda>y. f (x, y)) (B' x)) \<longlongrightarrow> g x) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3094 | (filtercomap (\<lambda>p. p x) (finite_subsets_at_top (B x)))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3095 | by (rule filterlim_compose[OF _ filterlim_filtercomap]) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3096 | qed auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3097 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3098 | show "\<forall>\<^sub>F h in H. sum (\<lambda>x. sum (\<lambda>y. f (x, y)) (h x)) Z \<in> X'" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3099 | unfolding H_def | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3100 | proof (subst eventually_INF_finite[OF \<open>finite Z\<close>], rule exI, safe) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3101 | fix x assume x: "x \<in> Z" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3102 | hence x': "x \<in> A" using 3 by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3103 | show "eventually (\<lambda>h. finite (h x) \<and> Y2 x \<subseteq> h x \<and> h x \<subseteq> B x) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3104 | (filtercomap (\<lambda>p. p x) (finite_subsets_at_top (B x)))" using 3 Y2[OF x'] | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3105 | by (intro eventually_filtercomapI) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3106 | (auto simp: eventually_finite_subsets_at_top intro: exI[of _ "Y2 x"]) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3107 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3108 | fix h | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3109 | assume *: "\<forall>x\<in>Z. finite (h x) \<and> Y2 x \<subseteq> h x \<and> h x \<subseteq> B x" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3110 | hence "sum (\<lambda>x. sum (\<lambda>y. f (x, y)) (h x)) Z = sum f (Sigma Z h)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3111 | using \<open>finite Z\<close> by (subst sum.Sigma) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3112 | also have "\<dots> \<in> X''" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3113 | using * 3 Y(1,2) by (intro Y; force simp: Y1_def Y2_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3114 | also have "X'' \<subseteq> X'" by fact | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3115 | finally show "sum (\<lambda>x. sum (\<lambda>y. f (x, y)) (h x)) Z \<in> X'" . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3116 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3117 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3118 |       have "H = (INF x\<in>SIGMA x:Z. {X. finite X \<and> X \<subseteq> B x}.
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3119 |                   principal {y. finite (y (fst x)) \<and> snd x \<subseteq> y (fst x) \<and> y (fst x) \<subseteq> B (fst x)})"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3120 | unfolding H_def finite_subsets_at_top_def filtercomap_INF filtercomap_principal | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3121 | by (simp add: INF_Sigma) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3122 | also have "\<dots> \<noteq> bot" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3123 | proof (rule INF_filter_not_bot, subst INF_principal_finite, goal_cases) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3124 | case (2 X) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3125 | define H' where | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3126 |           "H' = (\<Inter>x\<in>X. {y. finite (y (fst x)) \<and> snd x \<subseteq> y (fst x) \<and> y (fst x) \<subseteq> B (fst x)})"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3127 |         from 2 have "(\<lambda>x. \<Union>(y,Y)\<in>X. if x = y then Y else {}) \<in> H'"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3128 | by (force split: if_splits simp: H'_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3129 |         hence "H' \<noteq> {}" by blast
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3130 | thus "principal H' \<noteq> bot" by (simp add: principal_eq_bot_iff) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3131 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3132 | finally show "H \<noteq> bot" . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3133 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3134 | also have "X' \<subseteq> X" by fact | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3135 | finally show "sum g Z \<in> X" . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3136 | qed (insert Y(1,2), auto simp: Y1_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3137 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3138 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3139 | lemma has_sum_unique: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3140 |   fixes f :: "_ \<Rightarrow> 'a :: {topological_comm_monoid_add, t2_space}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3141 | assumes "(f has_sum x) A" "(f has_sum y) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3142 | shows "x = y" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3143 | using assms unfolding has_sum_def using tendsto_unique finite_subsets_at_top_neq_bot by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3144 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3145 | lemma has_sum_SigmaI: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3146 |   fixes f :: "_ \<Rightarrow> 'a :: {topological_comm_monoid_add, t3_space}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3147 | assumes f: "\<And>x. x \<in> A \<Longrightarrow> ((\<lambda>y. f (x, y)) has_sum g x) (B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3148 | assumes g: "(g has_sum S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3149 | assumes summable: "f summable_on Sigma A B" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3150 | shows "(f has_sum S) (Sigma A B)" | 
| 77388 | 3151 | by (metis f g has_sum_SigmaD has_sum_infsum has_sum_unique local.summable) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3152 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3153 | lemma summable_on_SigmaD1: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3154 |   fixes f :: "_ \<Rightarrow> _ \<Rightarrow> 'a :: {complete_uniform_space, uniform_topological_group_add, ab_group_add, topological_comm_monoid_add}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3155 | assumes f: "(\<lambda>(x,y). f x y) summable_on Sigma A B" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3156 | assumes x: "x \<in> A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3157 | shows "f x summable_on B x" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3158 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3159 |   have "(\<lambda>(x,y). f x y) summable_on Sigma {x} B"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3160 | using f by (rule summable_on_subset) (use x in auto) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3161 |   also have "?this \<longleftrightarrow> ((\<lambda>y. f x y) \<circ> snd) summable_on Sigma {x} B"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3162 | by (intro summable_on_cong) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3163 |   also have "\<dots> \<longleftrightarrow> (\<lambda>y. f x y) summable_on snd ` Sigma {x} B"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3164 | by (intro summable_on_reindex [symmetric] inj_onI) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3165 |   also have "snd ` Sigma {x} B = B x"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3166 | by (force simp: Sigma_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3167 | finally show ?thesis . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3168 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3169 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3170 | lemma has_sum_swap: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3171 | "(f has_sum S) (A \<times> B) \<longleftrightarrow> ((\<lambda>(x,y). f (y,x)) has_sum S) (B \<times> A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3172 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3173 | have "bij_betw (\<lambda>(x,y). (y,x)) (B \<times> A) (A \<times> B)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3174 | by (rule bij_betwI[of _ _ _ "\<lambda>(x,y). (y,x)"]) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3175 | from has_sum_reindex_bij_betw[OF this, where f = f] show ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3176 | by (simp add: case_prod_unfold) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3177 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3178 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3179 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3180 | lemma summable_on_swap: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3181 | "f summable_on (A \<times> B) \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) summable_on (B \<times> A)" | 
| 77388 | 3182 | by (metis has_sum_swap summable_on_def) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3183 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3184 | lemma has_sum_cmult_right_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3185 |   fixes c :: "'a :: {topological_semigroup_mult, field}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3186 | assumes "c \<noteq> 0" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3187 | shows "((\<lambda>x. c * f x) has_sum S) A \<longleftrightarrow> (f has_sum (S / c)) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3188 | using has_sum_cmult_right[of f A "S/c" c] | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3189 | has_sum_cmult_right[of "\<lambda>x. c * f x" A S "inverse c"] assms | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3190 | by (auto simp: field_simps) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3191 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3192 | lemma has_sum_cmult_left_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3193 |   fixes c :: "'a :: {topological_semigroup_mult, field}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3194 | assumes "c \<noteq> 0" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3195 | shows "((\<lambda>x. f x * c) has_sum S) A \<longleftrightarrow> (f has_sum (S / c)) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3196 | by (smt (verit, best) assms has_sum_cmult_right_iff has_sum_cong mult.commute) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3197 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3198 | lemma finite_nonzero_values_imp_summable_on: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3199 |   assumes "finite {x\<in>X. f x \<noteq> 0}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3200 | shows "f summable_on X" | 
| 77388 | 3201 | by (smt (verit, del_insts) Diff_iff assms mem_Collect_eq summable_on_cong_neutral summable_on_finite) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3202 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3203 | lemma summable_on_of_int_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3204 | "(\<lambda>x::'a. of_int (f x) :: 'b :: real_normed_algebra_1) summable_on A \<longleftrightarrow> f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3205 | proof | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3206 | assume "f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3207 | thus "(\<lambda>x. of_int (f x)) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3208 | by (rule summable_on_homomorphism) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3209 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3210 | assume "(\<lambda>x. of_int (f x) :: 'b) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3211 | then obtain S where "((\<lambda>x. of_int (f x) :: 'b) has_sum S) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3212 | by (auto simp: summable_on_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3213 | hence "(sum (\<lambda>x. of_int (f x) :: 'b) \<longlongrightarrow> S) (finite_subsets_at_top A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3214 | unfolding has_sum_def . | 
| 77388 | 3215 | moreover have "1/2 > (0 :: real)" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3216 | by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3217 | ultimately have "eventually (\<lambda>X. dist (sum (\<lambda>x. of_int (f x) :: 'b) X) S < 1/2) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3218 | (finite_subsets_at_top A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3219 | unfolding tendsto_iff by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3220 | then obtain X where X: "finite X" "X \<subseteq> A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3221 | "\<And>Y. finite Y \<Longrightarrow> X \<subseteq> Y \<Longrightarrow> Y \<subseteq> A \<Longrightarrow> dist (sum (\<lambda>x. of_int (f x)) Y) S < 1/2" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3222 | unfolding eventually_finite_subsets_at_top by metis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3223 | |
| 77388 | 3224 | have "sum f Y = sum f X" if "finite Y" "X \<subseteq> Y" "Y \<subseteq> A" for Y | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3225 | proof - | 
| 77388 | 3226 | have "dist (sum (\<lambda>x. of_int (f x)) X) S < 1/2" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3227 | by (intro X) auto | 
| 77388 | 3228 | moreover have "dist (sum (\<lambda>x. of_int (f x)) Y) S < 1/2" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3229 | by (intro X that) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3230 | ultimately have "dist (sum (\<lambda>x. of_int (f x)) X) (sum (\<lambda>x. of_int (f x) :: 'b) Y) < | 
| 77388 | 3231 | 1/2 + 1/2" | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3232 | using dist_triangle_less_add by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3233 | thus ?thesis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3234 | by (simp add: dist_norm flip: of_int_sum of_int_diff) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3235 | qed | 
| 77388 | 3236 |   then have "{x\<in>A. f x \<noteq> 0} \<subseteq> X"
 | 
| 3237 | by (smt (verit) X finite_insert insert_iff mem_Collect_eq subset_eq sum.insert) | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3238 |   with \<open>finite X\<close> have "finite {x\<in>A. f x \<noteq> 0}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3239 | using finite_subset by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3240 | thus "f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3241 | by (rule finite_nonzero_values_imp_summable_on) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3242 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3243 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3244 | lemma summable_on_of_nat_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3245 | "(\<lambda>x::'a. of_nat (f x) :: 'b :: real_normed_algebra_1) summable_on A \<longleftrightarrow> f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3246 | proof | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3247 | assume "f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3248 | thus "(\<lambda>x. of_nat (f x) :: 'b) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3249 | by (rule summable_on_homomorphism) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3250 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3251 | assume "(\<lambda>x. of_nat (f x) :: 'b) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3252 | hence "(\<lambda>x. of_int (int (f x)) :: 'b) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3253 | by simp | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3254 | also have "?this \<longleftrightarrow> (\<lambda>x. int (f x)) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3255 | by (rule summable_on_of_int_iff) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3256 | also have "\<dots> \<longleftrightarrow> f summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3257 | by (simp add: summable_on_discrete_iff) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3258 | finally show "f summable_on A" . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3259 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3260 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3261 | lemma infsum_of_nat: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3262 |   "infsum (\<lambda>x::'a. of_nat (f x) :: 'b :: {real_normed_algebra_1}) A = of_nat (infsum f A)"
 | 
| 77388 | 3263 | by (metis has_sum_infsum has_sum_of_nat infsumI infsum_def of_nat_0 summable_on_of_nat_iff) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3264 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3265 | lemma infsum_of_int: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3266 |   "infsum (\<lambda>x::'a. of_int (f x) :: 'b :: {real_normed_algebra_1}) A = of_int (infsum f A)"
 | 
| 77388 | 3267 | by (metis has_sum_infsum has_sum_of_int infsumI infsum_not_exists of_int_0 summable_on_of_int_iff) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3268 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3269 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3270 | lemma summable_on_SigmaI: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3271 |   fixes f :: "_ \<Rightarrow> 'a :: {linorder_topology, ordered_comm_monoid_add, topological_comm_monoid_add,
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3272 | conditionally_complete_linorder}" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3273 | assumes f: "\<And>x. x \<in> A \<Longrightarrow> ((\<lambda>y. f (x, y)) has_sum g x) (B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3274 | assumes g: "g summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3275 | assumes f_nonneg: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> B x \<Longrightarrow> f (x, y) \<ge> (0 :: 'a)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3276 | shows "f summable_on Sigma A B" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3277 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3278 | have g_nonneg: "g x \<ge> 0" if "x \<in> A" for x | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3279 | using f by (rule has_sum_nonneg) (use f_nonneg that in auto) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3280 | obtain C where C: "eventually (\<lambda>X. sum g X \<le> C) (finite_subsets_at_top A)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3281 | using summable_on_imp_bounded_partial_sums[OF g] by blast | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3282 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3283 | have sum_g_le: "sum g X \<le> C" if X: "finite X" "X \<subseteq> A" for X | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3284 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3285 | from C obtain X' where X': | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3286 | "finite X'" "X' \<subseteq> A" "\<And>Y. finite Y \<Longrightarrow> X' \<subseteq> Y \<Longrightarrow> Y \<subseteq> A \<Longrightarrow> sum g Y \<le> C" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3287 | unfolding eventually_finite_subsets_at_top by metis | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3288 | have "sum g X \<le> sum g (X \<union> X')" | 
| 77388 | 3289 | using X X' by (intro sum_mono2 g_nonneg) auto | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3290 | also have "\<dots> \<le> C" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3291 | using X X'(1,2) by (intro X'(3)) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3292 | finally show ?thesis . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3293 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3294 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3295 | have "sum f Y \<le> C" if Y: "finite Y" "Y \<subseteq> Sigma A B" for Y | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3296 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3297 |     define Y1 and Y2 where "Y1 = fst ` Y" and "Y2 = (\<lambda>x. snd ` {z\<in>Y. fst z = x})"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3298 | have Y12: "Y = Sigma Y1 Y2" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3299 | unfolding Y1_def Y2_def by force | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3300 | have [intro]: "finite Y1" "\<And>x. x \<in> Y1 \<Longrightarrow> finite (Y2 x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3301 | using Y unfolding Y1_def Y2_def by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3302 | have Y12_subset: "Y1 \<subseteq> A" "\<And>x. Y2 x \<subseteq> B x" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3303 | using Y by (auto simp: Y1_def Y2_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3304 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3305 | have "sum f Y = sum f (Sigma Y1 Y2)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3306 | by (simp add: Y12) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3307 | also have "\<dots> = (\<Sum>x\<in>Y1. \<Sum>y\<in>Y2 x. f (x, y))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3308 | by (subst sum.Sigma) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3309 | also have "\<dots> \<le> (\<Sum>x\<in>Y1. g x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3310 | proof (rule sum_mono) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3311 | fix x assume x: "x \<in> Y1" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3312 | show "(\<Sum>y\<in>Y2 x. f (x, y)) \<le> g x" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3313 | proof (rule has_sum_mono') | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3314 | show "((\<lambda>y. f (x, y)) has_sum (\<Sum>y\<in>Y2 x. f (x, y))) (Y2 x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3315 | using x by (intro has_sum_finite) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3316 | show "((\<lambda>y. f (x, y)) has_sum g x) (B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3317 | by (rule f) (use x Y12_subset in auto) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3318 | show "f (x, y) \<ge> 0" if "y \<in> B x - Y2 x" for y | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3319 | using x that Y12_subset by (intro f_nonneg) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3320 | qed (use Y12_subset in auto) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3321 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3322 | also have "\<dots> \<le> C" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3323 | using Y12_subset by (intro sum_g_le) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3324 | finally show ?thesis . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3325 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3326 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3327 | hence "\<forall>\<^sub>F X in finite_subsets_at_top (Sigma A B). sum f X \<le> C" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3328 | unfolding eventually_finite_subsets_at_top by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3329 | thus ?thesis | 
| 77388 | 3330 | by (metis SigmaE f_nonneg nonneg_bounded_partial_sums_imp_summable_on) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3331 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3332 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3333 | lemma summable_on_UnionI: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3334 |   fixes f :: "_ \<Rightarrow> 'a :: {linorder_topology, ordered_comm_monoid_add, topological_comm_monoid_add,
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3335 | conditionally_complete_linorder}" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3336 | assumes f: "\<And>x. x \<in> A \<Longrightarrow> (f has_sum g x) (B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3337 | assumes g: "g summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3338 | assumes f_nonneg: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> B x \<Longrightarrow> f y \<ge> (0 :: 'a)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3339 | assumes disj: "disjoint_family_on B A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3340 | shows "f summable_on (\<Union>x\<in>A. B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3341 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3342 | have "f \<circ> snd summable_on Sigma A B" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3343 | using assms by (intro summable_on_SigmaI[where g = g]) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3344 | also have "?this \<longleftrightarrow> f summable_on (snd ` Sigma A B)" using assms | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3345 | by (subst summable_on_reindex; force simp: disjoint_family_on_def inj_on_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3346 | also have "snd ` (Sigma A B) = (\<Union>x\<in>A. B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3347 | by force | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3348 | finally show ?thesis . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3349 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3350 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3351 | lemma summable_on_SigmaD: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3352 |   fixes f :: "'a \<times> 'b \<Rightarrow> 'c :: {topological_comm_monoid_add,t3_space}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3353 | assumes sum1: "f summable_on (Sigma A B)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3354 | assumes sum2: "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) summable_on (B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3355 | shows "(\<lambda>x. infsum (\<lambda>y. f (x, y)) (B x)) summable_on A" | 
| 77388 | 3356 | using assms unfolding summable_on_def | 
| 3357 | by (smt (verit, del_insts) assms has_sum_SigmaD has_sum_cong has_sum_infsum) | |
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3358 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3359 | lemma summable_on_UnionD: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3360 |   fixes f :: "'a \<Rightarrow> 'c :: {topological_comm_monoid_add,t3_space}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3361 | assumes sum1: "f summable_on (\<Union>x\<in>A. B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3362 | assumes sum2: "\<And>x. x \<in> A \<Longrightarrow> f summable_on (B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3363 | assumes disj: "disjoint_family_on B A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3364 | shows "(\<lambda>x. infsum f (B x)) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3365 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3366 | have "(\<Union>x\<in>A. B x) = snd ` Sigma A B" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3367 | by (force simp: Sigma_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3368 | with sum1 have "f summable_on (snd ` Sigma A B)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3369 | by simp | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3370 | also have "?this \<longleftrightarrow> (f \<circ> snd) summable_on (Sigma A B)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3371 | using disj by (intro summable_on_reindex inj_onI) (force simp: disjoint_family_on_def) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3372 | finally show "(\<lambda>x. infsum f (B x)) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3373 | using summable_on_SigmaD[of "f \<circ> snd" A B] sum2 by simp | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3374 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3375 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3376 | lemma summable_on_Union_iff: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3377 |   fixes f :: "_ \<Rightarrow> 'a :: {linorder_topology, ordered_comm_monoid_add, topological_comm_monoid_add,
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3378 | conditionally_complete_linorder, t3_space}" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3379 | assumes f: "\<And>x. x \<in> A \<Longrightarrow> (f has_sum g x) (B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3380 | assumes f_nonneg: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> B x \<Longrightarrow> f y \<ge> 0" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3381 | assumes disj: "disjoint_family_on B A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3382 | shows "f summable_on (\<Union>x\<in>A. B x) \<longleftrightarrow> g summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3383 | proof | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3384 | assume "g summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3385 | thus "f summable_on (\<Union>x\<in>A. B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3386 | using summable_on_UnionI[of A f B g] assms by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3387 | next | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3388 | assume "f summable_on (\<Union>x\<in>A. B x)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3389 | hence "(\<lambda>x. infsum f (B x)) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3390 | using assms by (intro summable_on_UnionD) (auto dest: has_sum_imp_summable) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3391 | also have "?this \<longleftrightarrow> g summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3392 | using assms by (intro summable_on_cong) (auto simp: infsumI) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3393 | finally show "g summable_on A" . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3394 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3395 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3396 | lemma has_sum_Sigma': | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3397 | fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3398 |     and f :: \<open>'a \<times> 'b \<Rightarrow> 'c::{comm_monoid_add,uniform_space,uniform_topological_group_add}\<close>
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3399 | assumes summableAB: "(f has_sum a) (Sigma A B)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3400 | assumes summableB: \<open>\<And>x. x\<in>A \<Longrightarrow> ((\<lambda>y. f (x, y)) has_sum (b x)) (B x)\<close> | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3401 | shows "(b has_sum a) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3402 | by (intro has_sum_Sigma[OF _ assms] uniformly_continuous_add) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3403 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3404 | lemma abs_summable_on_comparison_test': | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3405 | assumes "g summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3406 | assumes "\<And>x. x \<in> A \<Longrightarrow> norm (f x) \<le> g x" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3407 | shows "(\<lambda>x. norm (f x)) summable_on A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3408 | proof (rule Infinite_Sum.abs_summable_on_comparison_test) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3409 | have "g summable_on A \<longleftrightarrow> (\<lambda>x. norm (g x)) summable_on A" | 
| 77388 | 3410 | by (metis summable_on_iff_abs_summable_on_real) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3411 | with assms show "(\<lambda>x. norm (g x)) summable_on A" by blast | 
| 77388 | 3412 | qed (use assms in fastforce) | 
| 77357 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3413 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3414 | lemma has_sum_geometric_from_1: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3415 |   fixes z :: "'a :: {real_normed_field, banach}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3416 | assumes "norm z < 1" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3417 |   shows   "((\<lambda>n. z ^ n) has_sum (z / (1 - z))) {1..}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3418 | proof - | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3419 | have [simp]: "z \<noteq> 1" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3420 | using assms by auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3421 | have "(\<lambda>n. z ^ Suc n) sums (1 / (1 - z) - 1)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3422 | using geometric_sums[of z] assms by (subst sums_Suc_iff) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3423 | also have "1 / (1 - z) - 1 = z / (1 - z)" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3424 | by (auto simp: field_simps) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3425 | finally have "(\<lambda>n. z ^ Suc n) sums (z / (1 - z))" . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3426 | moreover have "summable (\<lambda>n. norm (z ^ Suc n))" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3427 | using assms | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3428 | by (subst summable_Suc_iff) (auto simp: norm_power intro!: summable_geometric) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3429 | ultimately have "((\<lambda>n. z ^ Suc n) has_sum (z / (1 - z))) UNIV" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3430 | by (intro norm_summable_imp_has_sum) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3431 | also have "?this \<longleftrightarrow> ?thesis" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3432 | by (intro has_sum_reindex_bij_witness[of _ "\<lambda>n. n-1" "\<lambda>n. n+1"]) auto | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3433 | finally show ?thesis . | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3434 | qed | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3435 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3436 | lemma has_sum_divide_const: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3437 |   fixes f :: "'a \<Rightarrow> 'b :: {topological_semigroup_mult, field, semiring_0}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3438 | shows "(f has_sum S) A \<Longrightarrow> ((\<lambda>x. f x / c) has_sum (S / c)) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3439 | using has_sum_cmult_right[of f A S "inverse c"] by (simp add: field_simps) | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3440 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3441 | lemma has_sum_uminusI: | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3442 |   fixes f :: "'a \<Rightarrow> 'b :: {topological_semigroup_mult, ring_1}"
 | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3443 | shows "(f has_sum S) A \<Longrightarrow> ((\<lambda>x. -f x) has_sum (-S)) A" | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3444 | using has_sum_cmult_right[of f A S "-1"] by simp | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3445 | |
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3446 | end | 
| 
e65d8ee80811
New material contributed by Manuel
 paulson <lp15@cam.ac.uk> parents: 
76999diff
changeset | 3447 |