author | paulson <lp15@cam.ac.uk> |
Tue, 25 Mar 2025 21:19:26 +0000 (5 weeks ago) | |
changeset 82349 | a854ca7ca7d9 |
parent 81150 | 3dd8035578b8 |
child 82522 | 62afd98e3f3e |
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:
74642
diff
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:
74642
diff
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:
76988
diff
changeset
|
25 |
theory Infinite_Sum |
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
wenzelm
parents:
76988
diff
changeset
|
26 |
imports |
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
wenzelm
parents:
76988
diff
changeset
|
27 |
Elementary_Topology |
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
wenzelm
parents:
76988
diff
changeset
|
28 |
"HOL-Library.Extended_Nonnegative_Real" |
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
wenzelm
parents:
76988
diff
changeset
|
29 |
"HOL-Library.Complex_Order" |
ff203584b36e
backed out changeset 7f7d5c93e36b: no longer required thanks to 9096703ed99e;
wenzelm
parents:
76988
diff
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 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
37 |
abbreviation has_sum (infixr \<open>has'_sum\<close> 46) where |
77359 | 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 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
40 |
definition summable_on :: "('a \<Rightarrow> 'b::{comm_monoid_add, topological_space}) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr \<open>summable'_on\<close> 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 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
46 |
abbreviation abs_summable_on :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr \<open>abs'_summable'_on\<close> 46) where |
74475
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) |
81097 | 50 |
"_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add" |
51 |
(\<open>(\<open>indent=3 notation=\<open>binder INFSUM\<close>\<close>INFSUM (_/:_)./ _)\<close> [0, 51, 10] 10) |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
52 |
syntax |
81097 | 53 |
"_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add" |
54 |
(\<open>(\<open>indent=2 notation=\<open>binder \<Sum>\<^sub>\<infinity>\<close>\<close>\<Sum>\<^sub>\<infinity>(_/\<in>_)./ _)\<close> [0, 51, 10] 10) |
|
80768 | 55 |
syntax_consts |
56 |
"_infsum" \<rightleftharpoons> infsum |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
57 |
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
|
58 |
"\<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
|
59 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
60 |
syntax (ASCII) |
81097 | 61 |
"_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(\<open>indent=3 notation=\<open>binder INFSUM\<close>\<close>INFSUM _./ _)\<close> [0, 10] 10) |
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
62 |
syntax |
81097 | 63 |
"_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(\<open>indent=2 notation=\<open>binder \<Sum>\<^sub>\<infinity>\<close>\<close>\<Sum>\<^sub>\<infinity>_./ _)\<close> [0, 10] 10) |
80768 | 64 |
syntax_consts |
65 |
"_univinfsum" \<rightleftharpoons> infsum |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
66 |
translations |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
67 |
"\<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
|
68 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
69 |
syntax (ASCII) |
81097 | 70 |
"_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(\<open>indent=3 notation=\<open>binder INFSUM\<close>\<close>INFSUM _ |/ _./ _)\<close> [0, 0, 10] 10) |
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
71 |
syntax |
81097 | 72 |
"_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(\<open>indent=2 notation=\<open>binder \<Sum>\<^sub>\<infinity>\<close>\<close>\<Sum>\<^sub>\<infinity>_ | (_)./ _)\<close> [0, 0, 10] 10) |
80768 | 73 |
syntax_consts |
74 |
"_qinfsum" \<rightleftharpoons> infsum |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
75 |
translations |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
76 |
"\<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
|
77 |
print_translation \<open> |
81150 | 78 |
[(\<^const_syntax>\<open>infsum\<close>, K (Collect_binder_tr' \<^syntax_const>\<open>_qinfsum\<close>))] |
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
79 |
\<close> |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
80 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
81 |
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
|
82 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
83 |
lemma infsumI: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
84 |
fixes f g :: \<open>'a \<Rightarrow> 'b::{comm_monoid_add, t2_space}\<close> |
77359 | 85 |
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
|
86 |
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
|
87 |
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
|
88 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
89 |
lemma infsum_eqI: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
90 |
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
|
91 |
assumes \<open>x = y\<close> |
77359 | 92 |
assumes \<open>(f has_sum x) A\<close> |
93 |
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
|
94 |
shows \<open>infsum f A = infsum g B\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
95 |
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
|
96 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
97 |
lemma infsum_eqI': |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
98 |
fixes f g :: \<open>'a \<Rightarrow> 'b::{comm_monoid_add, t2_space}\<close> |
77359 | 99 |
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
|
100 |
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
|
101 |
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
|
102 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
103 |
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
|
104 |
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
|
105 |
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
|
106 |
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
|
107 |
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
|
108 |
|
82349
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
109 |
lemma has_sum_unique: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
110 |
fixes f :: "_ \<Rightarrow> 'a :: {topological_comm_monoid_add, t2_space}" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
111 |
assumes "(f has_sum x) A" "(f has_sum y) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
112 |
shows "x = y" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
113 |
using assms infsumI by blast |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
114 |
|
77359 | 115 |
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:
74642
diff
changeset
|
116 |
using infsumI summable_on_def by blast |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
117 |
|
82349
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
118 |
lemma has_sum_iff: "(f has_sum S) A \<longleftrightarrow> f summable_on A \<and> infsum f A = S" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
119 |
using infsumI summable_iff_has_sum_infsum by blast |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
120 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
121 |
lemma has_sum_infsum[simp]: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
122 |
assumes \<open>f summable_on S\<close> |
77359 | 123 |
shows \<open>(f has_sum (infsum f S)) S\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
124 |
using assms summable_iff_has_sum_infsum by blast |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
125 |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
126 |
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
|
127 |
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
|
128 |
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
|
129 |
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
|
130 |
assumes \<open>\<And>x. x\<in>S\<inter>T \<Longrightarrow> f x = g x\<close> |
77359 | 131 |
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
|
132 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
133 |
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
|
134 |
= 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
|
135 |
proof |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
136 |
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
|
137 |
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
|
138 |
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
|
139 |
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
|
140 |
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
|
141 |
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
|
142 |
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
|
143 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
144 |
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:
75462
diff
changeset
|
145 |
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
|
146 |
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:
75462
diff
changeset
|
147 |
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:
74642
diff
changeset
|
148 |
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
|
149 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
150 |
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
|
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 |
next |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
153 |
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
|
154 |
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
|
155 |
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
|
156 |
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
|
157 |
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
|
158 |
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
|
159 |
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
|
160 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
161 |
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:
75462
diff
changeset
|
162 |
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
|
163 |
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:
75462
diff
changeset
|
164 |
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:
74642
diff
changeset
|
165 |
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
|
166 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
167 |
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
|
168 |
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
|
169 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
170 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
171 |
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
|
172 |
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
|
173 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
174 |
then show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
175 |
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
|
176 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
177 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
178 |
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
|
179 |
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
|
180 |
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
|
181 |
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
|
182 |
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
|
183 |
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
|
184 |
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
|
185 |
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
|
186 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
187 |
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
|
188 |
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
|
189 |
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
|
190 |
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
|
191 |
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
|
192 |
shows \<open>infsum f S = infsum g T\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
193 |
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
|
194 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
195 |
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
|
196 |
assumes "\<And>x. x\<in>A \<Longrightarrow> f x = g x" |
77359 | 197 |
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:
74642
diff
changeset
|
198 |
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
|
199 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
200 |
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
|
201 |
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
|
202 |
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
|
203 |
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
|
204 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
205 |
lemma infsum_cong: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
206 |
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
|
207 |
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
|
208 |
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
|
209 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
210 |
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
|
211 |
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
|
212 |
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
|
213 |
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
|
214 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
215 |
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
|
216 |
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
|
217 |
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
|
218 |
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
|
219 |
by auto |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
220 |
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
|
221 |
\<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
|
222 |
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
|
223 |
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
|
224 |
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
|
225 |
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
|
226 |
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
|
227 |
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
|
228 |
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
|
229 |
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
|
230 |
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
|
231 |
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
|
232 |
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
|
233 |
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
|
234 |
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
|
235 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
236 |
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
|
237 |
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
|
238 |
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
|
239 |
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
|
240 |
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
|
241 |
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
|
242 |
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
|
243 |
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
|
244 |
tendsto_compose_filtermap) |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
245 |
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
|
246 |
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
|
247 |
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
|
248 |
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
|
249 |
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
|
250 |
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
|
251 |
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
|
252 |
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
|
253 |
\<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
|
254 |
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
|
255 |
by simp |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
256 |
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
|
257 |
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
|
258 |
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
|
259 |
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
|
260 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
261 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
262 |
lemma |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
263 |
fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add}" |
77359 | 264 |
assumes \<open>(f has_sum b) B\<close> and \<open>(f has_sum a) A\<close> and AB: "A \<subseteq> B" |
265 |
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
|
266 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
267 |
have finite_subsets1: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
268 |
"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
|
269 |
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
|
270 |
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
|
271 |
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
|
272 |
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
|
273 |
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
|
274 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
275 |
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
|
276 |
by auto |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
277 |
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
|
278 |
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
|
279 |
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
|
280 |
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
|
281 |
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
|
282 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
283 |
have finite_subsets2: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
284 |
"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
|
285 |
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
|
286 |
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
|
287 |
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
|
288 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
289 |
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
|
290 |
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
|
291 |
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
|
292 |
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
|
293 |
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:
75462
diff
changeset
|
294 |
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
|
295 |
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
|
296 |
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
|
297 |
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
|
298 |
unfolding o_def |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
299 |
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
|
300 |
\<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
|
301 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
302 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
303 |
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
|
304 |
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
|
305 |
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
|
306 |
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
|
307 |
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
|
308 |
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
|
309 |
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
|
310 |
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
|
311 |
\<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
|
312 |
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
|
313 |
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
|
314 |
thus ?thesis |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
315 |
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
|
316 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
317 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
318 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
319 |
lemma |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
320 |
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
|
321 |
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
|
322 |
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
|
323 |
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
|
324 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
325 |
lemma |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
326 |
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
|
327 |
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
|
328 |
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:
74475
diff
changeset
|
329 |
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
|
330 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
331 |
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
|
332 |
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
|
333 |
(* Does this really require a linorder topology? (Instead of order topology.) *) |
77359 | 334 |
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
|
335 |
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
|
336 |
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
|
337 |
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
|
338 |
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
|
339 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
340 |
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
|
341 |
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
|
342 |
using assms(1,2) summable_on_def by auto |
77359 | 343 |
have \<open>(f' has_sum a) (A\<union>B)\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
344 |
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
|
345 |
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
|
346 |
by (meson has_sum_def) |
77359 | 347 |
have \<open>(g' has_sum b) (A\<union>B)\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
348 |
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
|
349 |
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
|
350 |
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
|
351 |
|
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
352 |
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
|
353 |
using assms by (auto simp: f'_def g'_def) |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
354 |
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:
75462
diff
changeset
|
355 |
by (intro eventually_finite_subsets_at_top_weakI sum_mono) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
356 |
then show ?thesis |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
357 |
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
|
358 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
359 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
360 |
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
|
361 |
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
|
362 |
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
|
363 |
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
|
364 |
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
|
365 |
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
|
366 |
shows "infsum f A \<le> infsum g B" |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
367 |
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
|
368 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
369 |
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
|
370 |
fixes f :: "'a\<Rightarrow>'b::{ordered_comm_monoid_add,linorder_topology}" |
77359 | 371 |
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
|
372 |
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
|
373 |
shows "x \<le> y" |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
374 |
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
|
375 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
376 |
lemma infsum_mono: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
377 |
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
|
378 |
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
|
379 |
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
|
380 |
shows "infsum f A \<le> infsum g A" |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
381 |
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
|
382 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
383 |
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
|
384 |
assumes "finite F" |
77359 | 385 |
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
|
386 |
using assms |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
387 |
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
|
388 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
389 |
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
|
390 |
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
|
391 |
assumes "finite F" |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
392 |
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
|
393 |
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
|
394 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
395 |
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
|
396 |
assumes "finite F" |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
397 |
shows "infsum f F = sum f F" |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
398 |
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
|
399 |
|
82349
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
400 |
lemma has_sum_finiteI: "finite A \<Longrightarrow> S = sum f A \<Longrightarrow> (f has_sum S) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
401 |
by simp |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
402 |
|
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
403 |
lemma has_sum_strict_mono_neutral: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
404 |
fixes f :: "'a \<Rightarrow> 'b :: {ordered_ab_group_add, topological_ab_group_add, linorder_topology}" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
405 |
assumes \<open>(f has_sum a) A\<close> and "(g has_sum b) B" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
406 |
assumes \<open>\<And>x. x \<in> A\<inter>B \<Longrightarrow> f x \<le> g x\<close> |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
407 |
assumes \<open>\<And>x. x \<in> A-B \<Longrightarrow> f x \<le> 0\<close> |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
408 |
assumes \<open>\<And>x. x \<in> B-A \<Longrightarrow> g x \<ge> 0\<close> |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
409 |
assumes \<open>x \<in> B\<close> \<open>if x \<in> A then f x < g x else 0 < g x\<close> |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
410 |
shows "a < b" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
411 |
proof - |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
412 |
define y where "y = (if x \<in> A then f x else 0)" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
413 |
have "a - y \<le> b - g x" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
414 |
proof (rule has_sum_mono_neutral) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
415 |
show "(f has_sum (a - y)) (A - (if x \<in> A then {x} else {}))" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
416 |
by (intro has_sum_Diff assms has_sum_finiteI) (auto simp: y_def) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
417 |
show "(g has_sum (b - g x)) (B - {x})" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
418 |
by (intro has_sum_Diff assms has_sum_finiteI) (use assms in auto) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
419 |
qed (use assms in \<open>auto split: if_splits\<close>) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
420 |
moreover have "y < g x" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
421 |
using assms(3,4,5)[of x] assms(6-) by (auto simp: y_def split: if_splits) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
422 |
ultimately show ?thesis |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
423 |
by (metis diff_strict_left_mono diff_strict_mono leD neqE) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
424 |
qed |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
425 |
|
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
426 |
lemma has_sum_strict_mono: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
427 |
fixes f :: "'a \<Rightarrow> 'b :: {ordered_ab_group_add, topological_ab_group_add, linorder_topology}" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
428 |
assumes \<open>(f has_sum a) A\<close> and "(g has_sum b) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
429 |
assumes \<open>\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x\<close> |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
430 |
assumes \<open>x \<in> A\<close> \<open>f x < g x\<close> |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
431 |
shows "a < b" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
432 |
by (rule has_sum_strict_mono_neutral[OF assms(1,2), where x = x]) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
433 |
(use assms(3-) in auto) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
434 |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
435 |
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
|
436 |
fixes f :: "'a \<Rightarrow> 'b::{comm_monoid_add,metric_space}" |
77359 | 437 |
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
|
438 |
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
|
439 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
440 |
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
|
441 |
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
|
442 |
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
|
443 |
using assms(2) by (rule tendstoD) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
444 |
thus ?thesis |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
445 |
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
|
446 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
447 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
448 |
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
|
449 |
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
|
450 |
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
|
451 |
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:
74642
diff
changeset
|
452 |
proof - |
77359 | 453 |
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:
74642
diff
changeset
|
454 |
by (simp add: summable_iff_has_sum_infsum) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
455 |
from this and \<open>\<epsilon> > 0\<close> show ?thesis |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
456 |
by (rule has_sum_finite_approximation) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
457 |
qed |
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
458 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
459 |
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
|
460 |
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
|
461 |
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
|
462 |
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
|
463 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
464 |
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
|
465 |
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
|
466 |
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
|
467 |
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
|
468 |
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
|
469 |
(\<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
|
470 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
471 |
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
|
472 |
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
|
473 |
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
|
474 |
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
|
475 |
using lim |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
476 |
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:
74642
diff
changeset
|
477 |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
478 |
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
|
479 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
480 |
from ev_P |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
481 |
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:
74475
diff
changeset
|
482 |
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
|
483 |
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
|
484 |
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
|
485 |
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
|
486 |
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
|
487 |
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
|
488 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
489 |
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:
74642
diff
changeset
|
490 |
proof - |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
491 |
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:
74642
diff
changeset
|
492 |
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:
74642
diff
changeset
|
493 |
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:
74642
diff
changeset
|
494 |
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:
74642
diff
changeset
|
495 |
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:
74642
diff
changeset
|
496 |
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:
74642
diff
changeset
|
497 |
also have \<open>\<dots> < 2 * d\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
498 |
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:
74642
diff
changeset
|
499 |
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:
74642
diff
changeset
|
500 |
qed |
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
501 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
502 |
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:
74642
diff
changeset
|
503 |
by (rule dist_triangle3) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
504 |
also have \<open>\<dots> < 2 * d + 2 * d\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
505 |
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:
74642
diff
changeset
|
506 |
also have \<open>\<dots> \<le> e\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
507 |
by (auto simp: d_def) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
508 |
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
|
509 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
510 |
then show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
511 |
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
|
512 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
513 |
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
|
514 |
by (simp add: cauchy_filter_metric_filtermap) |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
515 |
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
|
516 |
by (meson Cauchy_convergent UNIV_I complete_def convergent_def) |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
517 |
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:
75462
diff
changeset
|
518 |
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
|
519 |
then show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
520 |
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
|
521 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
522 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
523 |
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
|
524 |
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
|
525 |
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
|
526 |
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
|
527 |
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
|
528 |
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
|
529 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
530 |
(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
|
531 |
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
|
532 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
533 |
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
|
534 |
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
|
535 |
and A :: "'b set" |
77359 | 536 |
assumes "((\<lambda>x. norm (f x)) has_sum n) A" |
537 |
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
|
538 |
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
|
539 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
540 |
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
|
541 |
proof- |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
542 |
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
|
543 |
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
|
544 |
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
|
545 |
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
|
546 |
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
|
547 |
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
|
548 |
hence "norm a \<le> norm (sum f F) + \<epsilon>" |
74642
8af77cb50c6d
tuned proofs -- avoid z3, which is unavailable on arm64-linux;
wenzelm
parents:
74639
diff
changeset
|
549 |
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
|
550 |
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
|
551 |
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
|
552 |
also have "\<dots> \<le> n + \<epsilon>" |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
553 |
proof (intro add_right_mono [OF has_sum_mono_neutral]) |
77359 | 554 |
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:
75462
diff
changeset
|
555 |
by (simp add: \<open>finite F\<close>) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
556 |
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
|
557 |
finally show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
558 |
by assumption |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
559 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
560 |
thus ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
561 |
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
|
562 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
563 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
564 |
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
|
565 |
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
|
566 |
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
|
567 |
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
|
568 |
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
|
569 |
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
|
570 |
case True |
77359 | 571 |
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:
75462
diff
changeset
|
572 |
by (simp add: assms) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
573 |
then show ?thesis |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
574 |
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
|
575 |
next |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
576 |
case False |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
577 |
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
|
578 |
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
|
579 |
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
|
580 |
for X |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
581 |
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
|
582 |
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
|
583 |
proof(rule ccontr) |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
584 |
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
|
585 |
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
|
586 |
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
|
587 |
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
|
588 |
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
|
589 |
moreover have "open S" |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
590 |
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
|
591 |
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
|
592 |
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
|
593 |
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
|
594 |
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
|
595 |
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
|
596 |
by blast |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
597 |
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
|
598 |
unfolding S_def by auto |
74642
8af77cb50c6d
tuned proofs -- avoid z3, which is unavailable on arm64-linux;
wenzelm
parents:
74639
diff
changeset
|
599 |
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
|
600 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
601 |
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
|
602 |
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
|
603 |
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
|
604 |
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
|
605 |
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
|
606 |
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
|
607 |
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
|
608 |
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
|
609 |
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
|
610 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
611 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
612 |
lemma infsum_tendsto: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
613 |
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
|
614 |
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:
76999
diff
changeset
|
615 |
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
|
616 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
617 |
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
|
618 |
assumes \<open>\<And>x. x\<in>M \<Longrightarrow> f x = 0\<close> |
77359 | 619 |
shows \<open>(f has_sum 0) M\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
620 |
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
|
621 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
622 |
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
|
623 |
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
|
624 |
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
|
625 |
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
|
626 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
627 |
lemma infsum_0: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
628 |
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
|
629 |
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
|
630 |
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
|
631 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
632 |
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
|
633 |
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
|
634 |
by (simp_all add: infsum_0) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
635 |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
636 |
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
|
637 |
by (simp_all add: summable_on_0) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
638 |
|
77359 | 639 |
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
|
640 |
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
|
641 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
642 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
643 |
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
|
644 |
fixes f g :: "'a \<Rightarrow> 'b::{topological_comm_monoid_add}" |
77359 | 645 |
assumes \<open>(f has_sum a) A\<close> |
646 |
assumes \<open>(g has_sum b) A\<close> |
|
647 |
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
|
648 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
649 |
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
|
650 |
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
|
651 |
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
|
652 |
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
|
653 |
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
|
654 |
then show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
655 |
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
|
656 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
657 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
658 |
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
|
659 |
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
|
660 |
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
|
661 |
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
|
662 |
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:
76999
diff
changeset
|
663 |
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
|
664 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
665 |
lemma infsum_add: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
666 |
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
|
667 |
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
|
668 |
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
|
669 |
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
|
670 |
proof - |
77359 | 671 |
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:
76999
diff
changeset
|
672 |
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
|
673 |
then show ?thesis |
74639
f831b6e589dc
tuned proofs -- avoid z3, which is unavailable on arm64-linux;
wenzelm
parents:
74475
diff
changeset
|
674 |
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
|
675 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
676 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
677 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
678 |
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
|
679 |
fixes f :: "'a \<Rightarrow> 'b::topological_comm_monoid_add" |
77359 | 680 |
assumes "(f has_sum a) A" |
681 |
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
|
682 |
assumes disj: "A \<inter> B = {}" |
77359 | 683 |
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
|
684 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
685 |
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
|
686 |
and \<open>fB x = (if x \<notin> A then f x else 0)\<close> for x |
77359 | 687 |
have fA: \<open>(fA has_sum a) (A \<union> B)\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
688 |
by (smt (verit, ccfv_SIG) DiffD1 DiffD2 UnCI fA_def assms(1) has_sum_cong_neutral inf_sup_absorb) |
77359 | 689 |
have fB: \<open>(fB has_sum b) (A \<union> B)\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
690 |
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
|
691 |
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
|
692 |
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
|
693 |
show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
694 |
unfolding fAB |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
695 |
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
|
696 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
697 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
698 |
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
|
699 |
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
|
700 |
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
|
701 |
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
|
702 |
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
|
703 |
shows \<open>f summable_on (A \<union> B)\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
704 |
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
|
705 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
706 |
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
|
707 |
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
|
708 |
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
|
709 |
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
|
710 |
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
|
711 |
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:
74642
diff
changeset
|
712 |
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
|
713 |
|
75462
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
paulson <lp15@cam.ac.uk>
parents:
74791
diff
changeset
|
714 |
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:
74791
diff
changeset
|
715 |
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:
74791
diff
changeset
|
716 |
assumes "summable (\<lambda>n. norm (f n))" and "f sums S" |
77359 | 717 |
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:
74791
diff
changeset
|
718 |
unfolding has_sum_def tendsto_iff eventually_finite_subsets_at_top |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
719 |
proof clarsimp |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
720 |
fix \<epsilon>::real |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
721 |
assume "\<epsilon> > 0" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
722 |
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:
74791
diff
changeset
|
723 |
by (auto simp: summable_def) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
724 |
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:
74791
diff
changeset
|
725 |
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:
76999
diff
changeset
|
726 |
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:
76999
diff
changeset
|
727 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
728 |
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:
74791
diff
changeset
|
729 |
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:
74791
diff
changeset
|
730 |
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:
74791
diff
changeset
|
731 |
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:
74791
diff
changeset
|
732 |
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:
74791
diff
changeset
|
733 |
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:
74791
diff
changeset
|
734 |
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:
76999
diff
changeset
|
735 |
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:
74791
diff
changeset
|
736 |
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:
74791
diff
changeset
|
737 |
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:
74791
diff
changeset
|
738 |
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:
74791
diff
changeset
|
739 |
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:
74791
diff
changeset
|
740 |
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:
74791
diff
changeset
|
741 |
also have "\<dots> < \<epsilon>" by (rule N) auto |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
742 |
finally show ?thesis by (simp add: dist_norm norm_minus_commute) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
743 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
744 |
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:
76999
diff
changeset
|
745 |
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:
74791
diff
changeset
|
746 |
qed |
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
paulson <lp15@cam.ac.uk>
parents:
74791
diff
changeset
|
747 |
|
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
paulson <lp15@cam.ac.uk>
parents:
74791
diff
changeset
|
748 |
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:
74791
diff
changeset
|
749 |
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:
74791
diff
changeset
|
750 |
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:
74791
diff
changeset
|
751 |
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:
74791
diff
changeset
|
752 |
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:
74791
diff
changeset
|
753 |
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
|
754 |
|
82349
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
755 |
lemma sums_nonneg_imp_has_sum_strong: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
756 |
assumes "f sums (S::real)" "eventually (\<lambda>n. f n \<ge> 0) sequentially" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
757 |
shows "(f has_sum S) UNIV" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
758 |
proof - |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
759 |
from assms(2) obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> f n \<ge> 0" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
760 |
by (auto simp: eventually_at_top_linorder) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
761 |
from assms(1) have "summable f" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
762 |
by (simp add: sums_iff) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
763 |
hence "summable (\<lambda>n. f (n + N))" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
764 |
by (rule summable_ignore_initial_segment) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
765 |
hence "summable (\<lambda>n. norm (f (n + N)))" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
766 |
using N by simp |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
767 |
hence "summable (\<lambda>n. norm (f n))" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
768 |
using summable_iff_shift by blast |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
769 |
with assms(1) show ?thesis |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
770 |
using norm_summable_imp_has_sum by blast |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
771 |
qed |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
772 |
|
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
773 |
lemma sums_nonneg_imp_has_sum: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
774 |
assumes "f sums (S::real)" and "\<And>n. f n \<ge> 0" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
775 |
shows "(f has_sum S) UNIV" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
776 |
by (rule sums_nonneg_imp_has_sum_strong) (use assms in auto) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
777 |
|
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
778 |
lemma summable_nonneg_imp_summable_on_strong: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
779 |
assumes "summable f" "eventually (\<lambda>n. f n \<ge> (0::real)) sequentially" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
780 |
shows "f summable_on UNIV" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
781 |
using assms has_sum_iff sums_nonneg_imp_has_sum_strong by blast |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
782 |
|
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
783 |
lemma summable_nonneg_imp_summable_on: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
784 |
assumes "summable f" "\<And>n. f n \<ge> (0::real)" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
785 |
shows "f summable_on UNIV" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
786 |
by (rule summable_nonneg_imp_summable_on_strong) (use assms in auto) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
787 |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
788 |
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
|
789 |
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
|
790 |
\begin{itemize} |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
791 |
\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
|
792 |
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
|
793 |
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
|
794 |
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
|
795 |
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
|
796 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
797 |
\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
|
798 |
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
|
799 |
\end{itemize} |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
800 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
801 |
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
|
802 |
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
|
803 |
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
|
804 |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
805 |
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
|
806 |
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
|
807 |
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
|
808 |
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
|
809 |
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
|
810 |
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
|
811 |
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
|
812 |
proof - |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
813 |
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
|
814 |
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
|
815 |
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
|
816 |
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
|
817 |
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
|
818 |
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
|
819 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
820 |
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
|
821 |
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
|
822 |
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
|
823 |
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
|
824 |
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
|
825 |
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:
74642
diff
changeset
|
826 |
using plus_cont \<open>eventually E' uniformity\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
827 |
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:
74642
diff
changeset
|
828 |
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:
74642
diff
changeset
|
829 |
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:
74642
diff
changeset
|
830 |
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:
74642
diff
changeset
|
831 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
832 |
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
|
833 |
unfolding cauchy_filter_def le_filter_def by simp |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
834 |
then obtain P1 P2 |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
835 |
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:
74642
diff
changeset
|
836 |
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:
74642
diff
changeset
|
837 |
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
|
838 |
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
|
839 |
by auto |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
840 |
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
|
841 |
by (metis eventually_finite_subsets_at_top) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
842 |
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
|
843 |
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
|
844 |
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
|
845 |
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:
74642
diff
changeset
|
846 |
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:
74642
diff
changeset
|
847 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
848 |
have *: "E' (sum f F1', sum f F2')" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
849 |
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:
74642
diff
changeset
|
850 |
proof (intro DE'[where c = "sum f F0A"] P1P2E) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
851 |
have "P1 (sum f (F1' \<union> F0A))" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
852 |
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:
74642
diff
changeset
|
853 |
thus "P1 (sum f F1' + sum f F0A)" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
854 |
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:
74642
diff
changeset
|
855 |
next |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
856 |
have "P2 (sum f (F2' \<union> F0A))" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
857 |
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:
74642
diff
changeset
|
858 |
thus "P2 (sum f F2' + sum f F0A)" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
859 |
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:
74642
diff
changeset
|
860 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
861 |
|
77388 | 862 |
have "eventually (\<lambda>x. E' (x, sum f F0B)) (filtermap (sum f) (finite_subsets_at_top B))" |
863 |
and "eventually (\<lambda>x. E' (sum f F0B, x)) (filtermap (sum f) (finite_subsets_at_top B))" |
|
864 |
unfolding eventually_filtermap eventually_finite_subsets_at_top |
|
865 |
by (rule exI[of _ F0B]; use * in \<open>force simp: F0B_def\<close>)+ |
|
866 |
then |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
867 |
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:
74642
diff
changeset
|
868 |
unfolding eventually_prod_filter |
77388 | 869 |
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
|
870 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
871 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
872 |
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:
74642
diff
changeset
|
873 |
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:
74642
diff
changeset
|
874 |
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
|
875 |
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
|
876 |
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
|
877 |
then show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
878 |
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
|
879 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
880 |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
881 |
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
|
882 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
883 |
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
|
884 |
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
|
885 |
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
|
886 |
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
|
887 |
shows \<open>f summable_on B\<close> |
77388 | 888 |
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
|
889 |
|
77359 | 890 |
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
|
891 |
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
|
892 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
893 |
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
|
894 |
by auto |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
895 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
896 |
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
|
897 |
by simp |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
898 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
899 |
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
|
900 |
fixes f :: "'a \<Rightarrow> 'b::topological_comm_monoid_add" |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
901 |
assumes \<open>finite A\<close> |
77359 | 902 |
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:
76999
diff
changeset
|
903 |
assumes \<open>\<And>a a'. a\<in>A \<Longrightarrow> a'\<in>A \<Longrightarrow> a\<noteq>a' \<Longrightarrow> B a \<inter> B a' = {}\<close> |
77359 | 904 |
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:
76999
diff
changeset
|
905 |
using assms |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
906 |
proof (induction) |
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
907 |
case empty |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
908 |
then show ?case |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
909 |
by simp |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
910 |
next |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
911 |
case (insert x A) |
77359 | 912 |
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
|
913 |
by (simp add: insert.prems) |
77359 | 914 |
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
|
915 |
using insert by simp |
77359 | 916 |
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:
75462
diff
changeset
|
917 |
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
|
918 |
then show ?case |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
919 |
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
|
920 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
921 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
922 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
923 |
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
|
924 |
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
|
925 |
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
|
926 |
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
|
927 |
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
|
928 |
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:
76999
diff
changeset
|
929 |
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
|
930 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
931 |
lemma sum_infsum: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
932 |
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
|
933 |
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
|
934 |
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
|
935 |
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
|
936 |
shows \<open>sum (\<lambda>a. infsum f (B a)) A = infsum f (\<Union>a\<in>A. B a)\<close> |
77388 | 937 |
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
|
938 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
939 |
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
|
940 |
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
|
941 |
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
|
942 |
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
|
943 |
(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
|
944 |
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
|
945 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
946 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
947 |
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
|
948 |
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:
75462
diff
changeset
|
949 |
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
|
950 |
\<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
|
951 |
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
|
952 |
\<comment> \<open>For \<^class>\<open>t2_space\<close>, this is equivalent to \<open>isCont f x\<close> by @{thm [source] isCont_def}.\<close> |
77359 | 953 |
assumes infsum: \<open>(g has_sum x) S\<close> |
954 |
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
|
955 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
956 |
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
|
957 |
using infsum has_sum_def by blast |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
958 |
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:
75462
diff
changeset
|
959 |
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:
75462
diff
changeset
|
960 |
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:
75462
diff
changeset
|
961 |
using tendsto_cong f_sum |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
962 |
by (simp add: Lim_transform_eventually eventually_finite_subsets_at_top_weakI) |
77359 | 963 |
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
|
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 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
967 |
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
|
968 |
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:
75462
diff
changeset
|
969 |
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
|
970 |
\<comment> \<open>Not using \<^const>\<open>additive\<close> because it would add sort constraint \<^class>\<open>ab_group_add\<close>\<close> |
77359 | 971 |
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
|
972 |
\<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
|
973 |
assumes \<open>g summable_on S\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
974 |
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
|
975 |
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
|
976 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
977 |
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
|
978 |
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:
75462
diff
changeset
|
979 |
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
|
980 |
\<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
|
981 |
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
|
982 |
assumes \<open>g summable_on S\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
983 |
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:
74642
diff
changeset
|
984 |
using assms |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
985 |
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
|
986 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
987 |
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
|
988 |
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
|
989 |
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
|
990 |
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
|
991 |
\<comment> \<open>For \<^class>\<open>t2_space\<close>, this is equivalent to \<open>isCont f x\<close> by @{thm [source] isCont_def}.\<close> |
77359 | 992 |
assumes infsum: \<open>(g has_sum x) S\<close> |
993 |
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:
74642
diff
changeset
|
994 |
using assms |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
995 |
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
|
996 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
997 |
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
|
998 |
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
|
999 |
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
|
1000 |
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
|
1001 |
assumes \<open>g summable_on S\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1002 |
shows \<open>(f \<circ> g) summable_on S\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1003 |
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
|
1004 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1005 |
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
|
1006 |
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
|
1007 |
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
|
1008 |
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
|
1009 |
assumes \<open>g summable_on S\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1010 |
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
|
1011 |
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
|
1012 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1013 |
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
|
1014 |
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
|
1015 |
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
|
1016 |
assumes \<open>bdd_above (sum f ` {F. F\<subseteq>A \<and> finite F})\<close> |
77359 | 1017 |
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
|
1018 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1019 |
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
|
1020 |
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
|
1021 |
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
|
1022 |
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
|
1023 |
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:
76999
diff
changeset
|
1024 |
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:
76999
diff
changeset
|
1025 |
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:
76999
diff
changeset
|
1026 |
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:
76999
diff
changeset
|
1027 |
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
|
1028 |
next |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1029 |
fix a assume *: \<open>(SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F) < a\<close> |
77388 | 1030 |
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:
74642
diff
changeset
|
1031 |
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
|
1032 |
then show \<open>\<forall>\<^sub>F x in finite_subsets_at_top A. sum f x < a\<close> |
77388 | 1033 |
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
|
1034 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1035 |
then show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1036 |
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
|
1037 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1038 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1039 |
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
|
1040 |
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
|
1041 |
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
|
1042 |
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
|
1043 |
shows \<open>f summable_on A\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1044 |
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
|
1045 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1046 |
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
|
1047 |
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
|
1048 |
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
|
1049 |
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
|
1050 |
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:
74642
diff
changeset
|
1051 |
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
|
1052 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1053 |
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
|
1054 |
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
|
1055 |
assumes \<open>\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0\<close> |
77359 | 1056 |
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:
74642
diff
changeset
|
1057 |
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
|
1058 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1059 |
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
|
1060 |
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
|
1061 |
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
|
1062 |
shows \<open>f summable_on A\<close> |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1063 |
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
|
1064 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1065 |
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
|
1066 |
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
|
1067 |
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
|
1068 |
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:
74642
diff
changeset
|
1069 |
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
|
1070 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1071 |
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
|
1072 |
fixes f :: "'a \<Rightarrow> 'b::{ordered_comm_monoid_add,linorder_topology}" |
77359 | 1073 |
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
|
1074 |
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
|
1075 |
shows "a \<ge> 0" |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1076 |
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
|
1077 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1078 |
lemma infsum_nonneg: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1079 |
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:
74642
diff
changeset
|
1080 |
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
|
1081 |
shows "infsum f M \<ge> 0" (is "?lhs \<ge> _") |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1082 |
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:
74642
diff
changeset
|
1083 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1084 |
lemma has_sum_mono2: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1085 |
fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}" |
77359 | 1086 |
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:
74642
diff
changeset
|
1087 |
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:
74642
diff
changeset
|
1088 |
shows "S \<le> S'" |
77388 | 1089 |
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:
74642
diff
changeset
|
1090 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1091 |
lemma infsum_mono2: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1092 |
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:
74642
diff
changeset
|
1093 |
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:
74642
diff
changeset
|
1094 |
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:
74642
diff
changeset
|
1095 |
shows "infsum f A \<le> infsum f B" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1096 |
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:
74642
diff
changeset
|
1097 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1098 |
lemma finite_sum_le_has_sum: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1099 |
fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}" |
77359 | 1100 |
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:
74642
diff
changeset
|
1101 |
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:
74642
diff
changeset
|
1102 |
shows "sum f B \<le> S" |
77388 | 1103 |
by (meson assms has_sum_finite has_sum_mono2) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1104 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1105 |
lemma finite_sum_le_infsum: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1106 |
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:
74642
diff
changeset
|
1107 |
assumes "f summable_on A" "finite B" "B \<subseteq> A" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1108 |
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:
74642
diff
changeset
|
1109 |
shows "sum f B \<le> infsum f A" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1110 |
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
|
1111 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1112 |
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
|
1113 |
assumes \<open>inj_on h A\<close> |
77359 | 1114 |
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
|
1115 |
proof - |
77359 | 1116 |
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
|
1117 |
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
|
1118 |
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:
75462
diff
changeset
|
1119 |
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
|
1120 |
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:
75462
diff
changeset
|
1121 |
proof (intro tendsto_cong eventually_finite_subsets_at_top_weakI sum.reindex) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1122 |
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:
75462
diff
changeset
|
1123 |
using assms subset_inj_on by blast |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1124 |
qed |
77359 | 1125 |
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
|
1126 |
by (simp add: has_sum_def) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1127 |
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
|
1128 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1129 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1130 |
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
|
1131 |
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
|
1132 |
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
|
1133 |
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
|
1134 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1135 |
lemma infsum_reindex: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1136 |
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
|
1137 |
shows \<open>infsum g (h ` A) = infsum (g \<circ> h) A\<close> |
77388 | 1138 |
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
|
1139 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1140 |
lemma summable_on_reindex_bij_betw: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1141 |
assumes "bij_betw g A B" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1142 |
shows "(\<lambda>x. f (g x)) summable_on A \<longleftrightarrow> f summable_on B" |
77388 | 1143 |
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:
74642
diff
changeset
|
1144 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1145 |
lemma infsum_reindex_bij_betw: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1146 |
assumes "bij_betw g A B" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1147 |
shows "infsum (\<lambda>x. f (g x)) A = infsum f B" |
77388 | 1148 |
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
|
1149 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1150 |
lemma sum_uniformity: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1151 |
assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'b::{uniform_space,comm_monoid_add},y). x+y)\<close> |
77388 | 1152 |
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
|
1153 |
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
|
1154 |
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 | 1155 |
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
|
1156 |
case 0 |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1157 |
then show ?case |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1158 |
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
|
1159 |
next |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1160 |
case (Suc n) |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1161 |
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
|
1162 |
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
|
1163 |
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
|
1164 |
apply atomize_elim |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1165 |
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
|
1166 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1167 |
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
|
1168 |
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
|
1169 |
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
|
1170 |
by metis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1171 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1172 |
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
|
1173 |
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
|
1174 |
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
|
1175 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1176 |
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
|
1177 |
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
|
1178 |
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
|
1179 |
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
|
1180 |
case True |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1181 |
then show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1182 |
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
|
1183 |
next |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1184 |
case False |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1185 |
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
|
1186 |
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
|
1187 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1188 |
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
|
1189 |
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
|
1190 |
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
|
1191 |
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
|
1192 |
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
|
1193 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1194 |
from DM |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1195 |
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
|
1196 |
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
|
1197 |
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
|
1198 |
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
|
1199 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1200 |
with D2_N |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1201 |
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
|
1202 |
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
|
1203 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1204 |
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
|
1205 |
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
|
1206 |
qed |
77388 | 1207 |
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
|
1208 |
by auto |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1209 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1210 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1211 |
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
|
1212 |
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
|
1213 |
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
|
1214 |
assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'c,y). x+y)\<close> |
77359 | 1215 |
assumes summableAB: "(f has_sum a) (Sigma A B)" |
1216 |
assumes summableB: \<open>\<And>x. x\<in>A \<Longrightarrow> ((\<lambda>y. f (x, y)) has_sum b x) (B x)\<close> |
|
1217 |
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
|
1218 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1219 |
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
|
1220 |
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
|
1221 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1222 |
from summableB |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1223 |
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
|
1224 |
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
|
1225 |
from summableAB |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1226 |
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
|
1227 |
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
|
1228 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1229 |
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:
75462
diff
changeset
|
1230 |
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
|
1231 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1232 |
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
|
1233 |
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
|
1234 |
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
|
1235 |
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
|
1236 |
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
|
1237 |
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
|
1238 |
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
|
1239 |
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
|
1240 |
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
|
1241 |
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
|
1242 |
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
|
1243 |
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
|
1244 |
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
|
1245 |
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
|
1246 |
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
|
1247 |
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
|
1248 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1249 |
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
|
1250 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1251 |
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
|
1252 |
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
|
1253 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1254 |
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
|
1255 |
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:
75462
diff
changeset
|
1256 |
for M' :: \<open>'a set\<close> and g g' |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1257 |
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
|
1258 |
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
|
1259 |
by auto |
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 |
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
|
1262 |
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
|
1263 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1264 |
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
|
1265 |
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
|
1266 |
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:
74642
diff
changeset
|
1267 |
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
|
1268 |
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
|
1269 |
ultimately show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1270 |
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
|
1271 |
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
|
1272 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1273 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1274 |
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
|
1275 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1276 |
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
|
1277 |
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
|
1278 |
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
|
1279 |
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
|
1280 |
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
|
1281 |
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:
75462
diff
changeset
|
1282 |
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
|
1283 |
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:
75462
diff
changeset
|
1284 |
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
|
1285 |
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
|
1286 |
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
|
1287 |
with * show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1288 |
by auto |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1289 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1290 |
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
|
1291 |
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
|
1292 |
ultimately show ?thesis |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1293 |
unfolding FMB_def eventually_finite_subsets_at_top |
77388 | 1294 |
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
|
1295 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1296 |
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
|
1297 |
unfolding FMB_def eventually_finite_subsets_at_top |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1298 |
proof (rule exI[of _ G], safe) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1299 |
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:
74642
diff
changeset
|
1300 |
thus "D (\<Sum>(a,b)\<in>Y. f (a, b), a)" |
77388 | 1301 |
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:
74642
diff
changeset
|
1302 |
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
|
1303 |
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:
74642
diff
changeset
|
1304 |
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
|
1305 |
then show \<open>E (sum b M, a)\<close> |
77388 | 1306 |
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
|
1307 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1308 |
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
|
1309 |
using \<open>finite (fst ` G)\<close> and \<open>fst ` G \<subseteq> A\<close> |
77388 | 1310 |
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
|
1311 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1312 |
then show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1313 |
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
|
1314 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1315 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1316 |
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
|
1317 |
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
|
1318 |
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
|
1319 |
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
|
1320 |
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
|
1321 |
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
|
1322 |
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
|
1323 |
proof - |
77359 | 1324 |
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
|
1325 |
using has_sum_infsum by blast |
77359 | 1326 |
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
|
1327 |
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
|
1328 |
show ?thesis |
77388 | 1329 |
using plus_cont a b |
1330 |
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
|
1331 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1332 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1333 |
lemma infsum_Sigma: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1334 |
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
|
1335 |
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
|
1336 |
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
|
1337 |
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
|
1338 |
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
|
1339 |
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
|
1340 |
proof - |
77359 | 1341 |
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
|
1342 |
using has_sum_infsum by blast |
77359 | 1343 |
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
|
1344 |
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
|
1345 |
show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1346 |
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
|
1347 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1348 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1349 |
lemma infsum_Sigma': |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1350 |
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
|
1351 |
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
|
1352 |
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
|
1353 |
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
|
1354 |
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
|
1355 |
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
|
1356 |
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
|
1357 |
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
|
1358 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1359 |
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
|
1360 |
lemma |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1361 |
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
|
1362 |
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
|
1363 |
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
|
1364 |
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
|
1365 |
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
|
1366 |
proof - |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1367 |
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
|
1368 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1369 |
from assms |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1370 |
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
|
1371 |
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:
75462
diff
changeset
|
1372 |
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:
75462
diff
changeset
|
1373 |
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
|
1374 |
then show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1375 |
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
|
1376 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1377 |
show ?thesis1 |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1378 |
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
|
1379 |
show ?thesis2 |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1380 |
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
|
1381 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1382 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1383 |
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
|
1384 |
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
|
1385 |
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
|
1386 |
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
|
1387 |
shows \<open>infsum (\<lambda>x. infsum (\<lambda>y. f (x,y)) (B x)) A = infsum f (Sigma A B)\<close> |
77388 | 1388 |
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
|
1389 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1390 |
lemma infsum_swap: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1391 |
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
|
1392 |
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
|
1393 |
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
|
1394 |
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
|
1395 |
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
|
1396 |
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
|
1397 |
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
|
1398 |
proof - |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1399 |
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:
75462
diff
changeset
|
1400 |
by (simp add: assms(2) summable_on_cong) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1401 |
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:
75462
diff
changeset
|
1402 |
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
|
1403 |
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:
75462
diff
changeset
|
1404 |
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
|
1405 |
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
|
1406 |
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
|
1407 |
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
|
1408 |
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
|
1409 |
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:
75462
diff
changeset
|
1410 |
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:
74642
diff
changeset
|
1411 |
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
|
1412 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1413 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1414 |
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
|
1415 |
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
|
1416 |
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
|
1417 |
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
|
1418 |
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
|
1419 |
proof - |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1420 |
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:
75462
diff
changeset
|
1421 |
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
|
1422 |
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:
75462
diff
changeset
|
1423 |
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
|
1424 |
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
|
1425 |
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
|
1426 |
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
|
1427 |
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
|
1428 |
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:
75462
diff
changeset
|
1429 |
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:
74642
diff
changeset
|
1430 |
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
|
1431 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1432 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1433 |
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
|
1434 |
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
|
1435 |
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
|
1436 |
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
|
1437 |
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
|
1438 |
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
|
1439 |
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
|
1440 |
proof (rule ccontr) |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1441 |
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
|
1442 |
have ex: \<open>f summable_on (A-{x})\<close> |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1443 |
by (rule summable_on_cofin_subset) (use assms in auto) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1444 |
have pos: \<open>infsum f (A - {x}) \<ge> 0\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1445 |
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
|
1446 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1447 |
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
|
1448 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1449 |
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:
74642
diff
changeset
|
1450 |
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
|
1451 |
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:
74642
diff
changeset
|
1452 |
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
|
1453 |
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:
74642
diff
changeset
|
1454 |
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
|
1455 |
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
|
1456 |
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
|
1457 |
finally show False |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1458 |
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
|
1459 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1460 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1461 |
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
|
1462 |
fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add,ordered_ab_group_add,linorder_topology}" |
77359 | 1463 |
assumes "(f has_sum a) A" \<open>a \<le> 0\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1464 |
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
|
1465 |
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
|
1466 |
shows "f x = 0" |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1467 |
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
|
1468 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1469 |
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
|
1470 |
fixes f :: "'a \<Rightarrow> 'b :: {topological_semigroup_mult, semiring_0}" |
77359 | 1471 |
assumes \<open>(f has_sum a) A\<close> |
1472 |
shows "((\<lambda>x. f x * c) has_sum (a * c)) A" |
|
77388 | 1473 |
using assms tendsto_mult_right |
1474 |
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
|
1475 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1476 |
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
|
1477 |
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
|
1478 |
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
|
1479 |
shows "infsum (\<lambda>x. f x * c) A = infsum f A * c" |
77388 | 1480 |
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
|
1481 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1482 |
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
|
1483 |
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
|
1484 |
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
|
1485 |
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
|
1486 |
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
|
1487 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1488 |
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
|
1489 |
fixes f :: "'a \<Rightarrow> 'b :: {topological_semigroup_mult, semiring_0}" |
77359 | 1490 |
assumes \<open>(f has_sum a) A\<close> |
1491 |
shows "((\<lambda>x. c * f x) has_sum (c * a)) A" |
|
77388 | 1492 |
using assms tendsto_mult_left |
1493 |
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
|
1494 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1495 |
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
|
1496 |
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
|
1497 |
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
|
1498 |
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:
75462
diff
changeset
|
1499 |
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
|
1500 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1501 |
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
|
1502 |
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
|
1503 |
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
|
1504 |
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
|
1505 |
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
|
1506 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1507 |
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
|
1508 |
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
|
1509 |
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
|
1510 |
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
|
1511 |
proof |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1512 |
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
|
1513 |
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
|
1514 |
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
|
1515 |
next |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1516 |
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
|
1517 |
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
|
1518 |
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
|
1519 |
then show \<open>f summable_on A\<close> |
77388 | 1520 |
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
|
1521 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1522 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1523 |
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
|
1524 |
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
|
1525 |
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
|
1526 |
shows "(\<lambda>x. c * f x) summable_on A \<longleftrightarrow> f summable_on A" |
77388 | 1527 |
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
|
1528 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1529 |
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
|
1530 |
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
|
1531 |
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:
75462
diff
changeset
|
1532 |
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
|
1533 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1534 |
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
|
1535 |
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
|
1536 |
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:
75462
diff
changeset
|
1537 |
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
|
1538 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1539 |
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
|
1540 |
assumes \<open>finite F\<close> |
77359 | 1541 |
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
|
1542 |
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
|
1543 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1544 |
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
|
1545 |
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
|
1546 |
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:
75462
diff
changeset
|
1547 |
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
|
1548 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1549 |
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
|
1550 |
\<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
|
1551 |
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
|
1552 |
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
|
1553 |
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
|
1554 |
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
|
1555 |
proof (rule notI) |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1556 |
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
|
1557 |
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
|
1558 |
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
|
1559 |
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
|
1560 |
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
|
1561 |
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
|
1562 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1563 |
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
|
1564 |
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
|
1565 |
from assms |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1566 |
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
|
1567 |
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
|
1568 |
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
|
1569 |
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
|
1570 |
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
|
1571 |
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
|
1572 |
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
|
1573 |
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:
74642
diff
changeset
|
1574 |
finally show ?thesis . |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1575 |
qed |
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1576 |
then show False |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1577 |
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
|
1578 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1579 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1580 |
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
|
1581 |
\<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
|
1582 |
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
|
1583 |
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
|
1584 |
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:
75462
diff
changeset
|
1585 |
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
|
1586 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1587 |
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
|
1588 |
fixes f :: \<open>'a \<Rightarrow> 'b::topological_ab_group_add\<close> |
77359 | 1589 |
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
|
1590 |
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
|
1591 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1592 |
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
|
1593 |
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
|
1594 |
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
|
1595 |
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
|
1596 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1597 |
lemma infsum_uminus: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
1598 |
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
|
1599 |
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
|
1600 |
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
|
1601 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1602 |
lemma has_sum_le_finite_sums: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1603 |
fixes a :: \<open>'a::{comm_monoid_add,topological_space,linorder_topology}\<close> |
77359 | 1604 |
assumes \<open>(f has_sum a) A\<close> |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1605 |
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:
74642
diff
changeset
|
1606 |
shows \<open>a \<le> b\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1607 |
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:
74642
diff
changeset
|
1608 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1609 |
lemma infsum_le_finite_sums: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1610 |
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:
74642
diff
changeset
|
1611 |
assumes \<open>f summable_on A\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1612 |
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:
74642
diff
changeset
|
1613 |
shows \<open>infsum f A \<le> b\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1614 |
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:
74642
diff
changeset
|
1615 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1616 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1617 |
lemma summable_on_scaleR_left [intro]: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1618 |
fixes c :: \<open>'a :: real_normed_vector\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1619 |
assumes "c \<noteq> 0 \<Longrightarrow> f summable_on A" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1620 |
shows "(\<lambda>x. f x *\<^sub>R c) summable_on A" |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1621 |
proof (cases \<open>c = 0\<close>) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1622 |
case False |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1623 |
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:
75462
diff
changeset
|
1624 |
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:
75462
diff
changeset
|
1625 |
then show ?thesis |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1626 |
by (metis (mono_tags, lifting) comp_apply summable_on_cong) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1627 |
qed auto |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1628 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1629 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1630 |
lemma summable_on_scaleR_right [intro]: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1631 |
fixes f :: \<open>'a \<Rightarrow> 'b :: real_normed_vector\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1632 |
assumes "c \<noteq> 0 \<Longrightarrow> f summable_on A" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1633 |
shows "(\<lambda>x. c *\<^sub>R f x) summable_on A" |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1634 |
proof (cases \<open>c = 0\<close>) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1635 |
case False |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1636 |
then have "(*\<^sub>R) c \<circ> f summable_on A" |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1637 |
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:
75462
diff
changeset
|
1638 |
then show ?thesis |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1639 |
by (metis (mono_tags, lifting) comp_apply summable_on_cong) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1640 |
qed auto |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1641 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1642 |
lemma infsum_scaleR_left: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1643 |
fixes c :: \<open>'a :: real_normed_vector\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1644 |
assumes "c \<noteq> 0 \<Longrightarrow> f summable_on A" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1645 |
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:
75462
diff
changeset
|
1646 |
proof (cases \<open>c = 0\<close>) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1647 |
case False |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1648 |
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:
75462
diff
changeset
|
1649 |
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:
75462
diff
changeset
|
1650 |
then show ?thesis |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1651 |
by (metis (mono_tags, lifting) comp_apply infsum_cong) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1652 |
qed auto |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1653 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1654 |
lemma infsum_scaleR_right: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1655 |
fixes f :: \<open>'a \<Rightarrow> 'b :: real_normed_vector\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1656 |
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:
74642
diff
changeset
|
1657 |
proof - |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1658 |
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:
74642
diff
changeset
|
1659 |
by auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1660 |
then show ?thesis |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1661 |
proof cases |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1662 |
case summable |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1663 |
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:
75462
diff
changeset
|
1664 |
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:
74642
diff
changeset
|
1665 |
then show ?thesis |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1666 |
by (metis (mono_tags, lifting) comp_apply infsum_cong) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1667 |
next |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1668 |
case c0 |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1669 |
then show ?thesis by auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1670 |
next |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1671 |
case not_summable |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1672 |
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:
74642
diff
changeset
|
1673 |
proof (rule notI) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1674 |
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:
74642
diff
changeset
|
1675 |
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:
74642
diff
changeset
|
1676 |
using summable_on_scaleR_right by blast |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1677 |
with not_summable show False |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1678 |
by simp |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1679 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1680 |
then show ?thesis |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1681 |
by (simp add: infsum_not_exists not_summable(1)) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1682 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1683 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1684 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1685 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1686 |
lemma infsum_Un_Int: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1687 |
fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add, t2_space}" |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1688 |
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:
74642
diff
changeset
|
1689 |
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:
74642
diff
changeset
|
1690 |
proof - |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1691 |
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:
76999
diff
changeset
|
1692 |
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:
76999
diff
changeset
|
1693 |
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:
76999
diff
changeset
|
1694 |
using assms(2) infsum_Un_disjoint by fastforce |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1695 |
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:
76999
diff
changeset
|
1696 |
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:
74642
diff
changeset
|
1697 |
ultimately show ?thesis |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1698 |
by auto |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1699 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1700 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1701 |
lemma inj_combinator': |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1702 |
assumes "x \<notin> F" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1703 |
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:
74642
diff
changeset
|
1704 |
proof - |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1705 |
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:
74642
diff
changeset
|
1706 |
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:
74642
diff
changeset
|
1707 |
thus ?thesis |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1708 |
by (simp add: o_def) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1709 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1710 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1711 |
lemma infsum_prod_PiE: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1712 |
\<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:
74642
diff
changeset
|
1713 |
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:
74642
diff
changeset
|
1714 |
assumes finite: "finite A" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1715 |
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:
74642
diff
changeset
|
1716 |
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:
74642
diff
changeset
|
1717 |
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:
74642
diff
changeset
|
1718 |
proof (use finite assms(2-) in induction) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1719 |
case empty |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1720 |
then show ?case |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1721 |
by auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1722 |
next |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1723 |
case (insert x F) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1724 |
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:
74642
diff
changeset
|
1725 |
unfolding PiE_insert_eq |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1726 |
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:
74642
diff
changeset
|
1727 |
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:
74642
diff
changeset
|
1728 |
by (rule prod.cong) (use insert.hyps in auto) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1729 |
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:
74642
diff
changeset
|
1730 |
using \<open>x \<notin> F\<close> by (rule inj_combinator') |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1731 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1732 |
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:
74642
diff
changeset
|
1733 |
using insert.prems(2) . |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1734 |
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:
74642
diff
changeset
|
1735 |
by (simp only: pi) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1736 |
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:
74642
diff
changeset
|
1737 |
((\<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:
74642
diff
changeset
|
1738 |
using inj by (rule summable_on_reindex) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1739 |
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:
74642
diff
changeset
|
1740 |
using insert.hyps by (intro prod.cong) auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1741 |
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:
74642
diff
changeset
|
1742 |
(\<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:
74642
diff
changeset
|
1743 |
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:
74642
diff
changeset
|
1744 |
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:
74642
diff
changeset
|
1745 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1746 |
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:
74642
diff
changeset
|
1747 |
by (rule summable_on_Sigma_banach) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1748 |
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:
75462
diff
changeset
|
1749 |
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:
74642
diff
changeset
|
1750 |
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:
75462
diff
changeset
|
1751 |
using summable_on_cmult_right' that by blast |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1752 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1753 |
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:
74642
diff
changeset
|
1754 |
= (\<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:
75462
diff
changeset
|
1755 |
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:
74642
diff
changeset
|
1756 |
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:
75462
diff
changeset
|
1757 |
using insert.hyps by auto |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1758 |
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:
75462
diff
changeset
|
1759 |
using prod by presburger |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1760 |
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:
75462
diff
changeset
|
1761 |
using infsum_Sigma'_banach summable2 by force |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1762 |
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:
75462
diff
changeset
|
1763 |
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:
74642
diff
changeset
|
1764 |
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:
75462
diff
changeset
|
1765 |
using insert summable3 by auto |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1766 |
finally show ?case |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1767 |
by simp |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1768 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1769 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1770 |
lemma infsum_prod_PiE_abs: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1771 |
\<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:
74642
diff
changeset
|
1772 |
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:
74642
diff
changeset
|
1773 |
assumes finite: "finite A" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1774 |
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:
74642
diff
changeset
|
1775 |
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:
74642
diff
changeset
|
1776 |
proof (use finite assms(2) in induction) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1777 |
case empty |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1778 |
then show ?case |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1779 |
by auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1780 |
next |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1781 |
case (insert x A) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1782 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1783 |
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:
74642
diff
changeset
|
1784 |
unfolding PiE_insert_eq |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1785 |
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:
75462
diff
changeset
|
1786 |
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:
74642
diff
changeset
|
1787 |
by (rule prod.cong) (use insert.hyps in auto) |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1788 |
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:
75462
diff
changeset
|
1789 |
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:
74642
diff
changeset
|
1790 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1791 |
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:
74642
diff
changeset
|
1792 |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1793 |
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:
74642
diff
changeset
|
1794 |
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:
74642
diff
changeset
|
1795 |
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:
74642
diff
changeset
|
1796 |
proof - |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1797 |
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:
75462
diff
changeset
|
1798 |
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:
74642
diff
changeset
|
1799 |
using that by (auto simp: B'_def) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1800 |
have [simp]: \<open>finite (Pi\<^sub>E F B')\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1801 |
by (simp add: finite_PiE) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1802 |
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:
74642
diff
changeset
|
1803 |
using that by (auto simp: B'_def) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1804 |
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:
74642
diff
changeset
|
1805 |
unfolding B'_def using P that |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1806 |
by auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1807 |
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:
75462
diff
changeset
|
1808 |
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:
74642
diff
changeset
|
1809 |
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:
75462
diff
changeset
|
1810 |
by (simp add: sum_mono2) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1811 |
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:
75462
diff
changeset
|
1812 |
by (simp add: prod_norm) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1813 |
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:
74642
diff
changeset
|
1814 |
proof (use \<open>finite F\<close> in induction) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1815 |
case empty |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1816 |
then show ?case by simp |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1817 |
next |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1818 |
case (insert x F) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1819 |
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:
75462
diff
changeset
|
1820 |
by (simp add: inj_combinator' insert.hyps) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1821 |
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:
74642
diff
changeset
|
1822 |
= (\<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:
75462
diff
changeset
|
1823 |
by (simp add: pi sum.reindex case_prod_unfold) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1824 |
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:
76999
diff
changeset
|
1825 |
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:
74642
diff
changeset
|
1826 |
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:
75462
diff
changeset
|
1827 |
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:
74642
diff
changeset
|
1828 |
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:
76999
diff
changeset
|
1829 |
using insert by force |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1830 |
finally show ?case . |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1831 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1832 |
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:
75462
diff
changeset
|
1833 |
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:
74642
diff
changeset
|
1834 |
finally show ?thesis . |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1835 |
qed |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1836 |
then have "bdd_above |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1837 |
(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:
76999
diff
changeset
|
1838 |
using insert.hyps insert.prems by (intro bdd_aboveI) blast |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1839 |
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:
75462
diff
changeset
|
1840 |
using nonneg_bdd_above_summable_on |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1841 |
by (metis (mono_tags, lifting) Collect_cong norm_ge_zero) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1842 |
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:
74642
diff
changeset
|
1843 |
by (simp only: pi) |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1844 |
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:
75462
diff
changeset
|
1845 |
((\<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:
74642
diff
changeset
|
1846 |
using inj by (subst summable_on_reindex) (auto simp: o_def) |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1847 |
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:
74642
diff
changeset
|
1848 |
using insert.hyps by (intro prod.cong) auto |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1849 |
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:
75462
diff
changeset
|
1850 |
(\<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:
74642
diff
changeset
|
1851 |
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:
75462
diff
changeset
|
1852 |
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:
74642
diff
changeset
|
1853 |
|
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1854 |
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:
75462
diff
changeset
|
1855 |
= (\<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:
75462
diff
changeset
|
1856 |
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:
75462
diff
changeset
|
1857 |
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:
76999
diff
changeset
|
1858 |
using prod insert.hyps by auto |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1859 |
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:
75462
diff
changeset
|
1860 |
using abs_summable_summable infsum_Sigma'_banach summable2 by fastforce |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1861 |
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:
75462
diff
changeset
|
1862 |
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:
76999
diff
changeset
|
1863 |
finally show ?case |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1864 |
by (simp add: insert) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1865 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1866 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1867 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1868 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1869 |
subsection \<open>Absolute convergence\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1870 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1871 |
lemma abs_summable_countable: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1872 |
assumes \<open>f abs_summable_on A\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1873 |
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:
74642
diff
changeset
|
1874 |
proof - |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1875 |
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:
74642
diff
changeset
|
1876 |
proof (rule ccontr) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1877 |
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:
74642
diff
changeset
|
1878 |
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:
74642
diff
changeset
|
1879 |
proof - |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1880 |
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:
74642
diff
changeset
|
1881 |
by (meson real_arch_simple) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1882 |
from * |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1883 |
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:
74642
diff
changeset
|
1884 |
by (meson finite_if_finite_subsets_card_bdd nle_le) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1885 |
have \<open>b \<le> of_nat b' * t\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1886 |
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:
74642
diff
changeset
|
1887 |
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:
74642
diff
changeset
|
1888 |
by (simp add: cardF that) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1889 |
also have \<open>\<dots> = sum (\<lambda>x. t) F\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1890 |
by simp |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1891 |
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:
74642
diff
changeset
|
1892 |
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:
74642
diff
changeset
|
1893 |
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:
74642
diff
changeset
|
1894 |
using \<open>finite F\<close> by (rule infsum_finite[symmetric]) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1895 |
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:
74642
diff
changeset
|
1896 |
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:
74642
diff
changeset
|
1897 |
finally show ?thesis . |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1898 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1899 |
then show False |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1900 |
by (meson gt_ex linorder_not_less) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1901 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1902 |
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:
74642
diff
changeset
|
1903 |
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:
74642
diff
changeset
|
1904 |
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:
74642
diff
changeset
|
1905 |
proof safe |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1906 |
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:
74642
diff
changeset
|
1907 |
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:
74642
diff
changeset
|
1908 |
have "i \<ge> 1" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1909 |
by (simp add: i_def) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1910 |
moreover have "real i \<ge> 1 / norm (f x)" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1911 |
unfolding i_def by linarith |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1912 |
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:
74642
diff
changeset
|
1913 |
by (auto simp: divide_simps mult_ac) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1914 |
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:
74642
diff
changeset
|
1915 |
using \<open>x \<in> A\<close> by auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1916 |
qed auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1917 |
finally show ?thesis . |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1918 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1919 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1920 |
(* 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:
74642
diff
changeset
|
1921 |
lemma summable_on_iff_abs_summable_on_real: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1922 |
fixes f :: \<open>'a \<Rightarrow> real\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1923 |
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:
74642
diff
changeset
|
1924 |
proof (rule iffI) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1925 |
assume \<open>f summable_on A\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1926 |
define n A\<^sub>p A\<^sub>n |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
1927 |
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:
75462
diff
changeset
|
1928 |
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:
74642
diff
changeset
|
1929 |
by (auto simp: A\<^sub>p_def A\<^sub>n_def) |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1930 |
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:
74642
diff
changeset
|
1931 |
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:
75462
diff
changeset
|
1932 |
then have \<open>n summable_on A\<^sub>p\<close> |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1933 |
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:
75462
diff
changeset
|
1934 |
moreover have \<open>n summable_on A\<^sub>n\<close> |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1935 |
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:
75462
diff
changeset
|
1936 |
ultimately show \<open>n summable_on A\<close> |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1937 |
using A summable_on_Un_disjoint by blast |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1938 |
next |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1939 |
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:
74642
diff
changeset
|
1940 |
using abs_summable_summable by blast |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1941 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1942 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1943 |
lemma abs_summable_on_Sigma_iff: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1944 |
shows "f abs_summable_on Sigma A B \<longleftrightarrow> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1945 |
(\<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:
74642
diff
changeset
|
1946 |
((\<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:
74642
diff
changeset
|
1947 |
proof (intro iffI conjI ballI) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1948 |
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:
74642
diff
changeset
|
1949 |
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:
75462
diff
changeset
|
1950 |
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:
74642
diff
changeset
|
1951 |
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:
74642
diff
changeset
|
1952 |
using summable_on_iff_abs_summable_on_real by force |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1953 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1954 |
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:
74642
diff
changeset
|
1955 |
proof - |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1956 |
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:
75462
diff
changeset
|
1957 |
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:
74642
diff
changeset
|
1958 |
then show ?thesis |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
1959 |
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:
74642
diff
changeset
|
1960 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1961 |
next |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1962 |
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:
74642
diff
changeset
|
1963 |
(\<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:
74642
diff
changeset
|
1964 |
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:
74642
diff
changeset
|
1965 |
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:
74642
diff
changeset
|
1966 |
proof - |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1967 |
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:
74642
diff
changeset
|
1968 |
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:
74642
diff
changeset
|
1969 |
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:
74642
diff
changeset
|
1970 |
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:
74642
diff
changeset
|
1971 |
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:
75462
diff
changeset
|
1972 |
by (simp add: sum.Sigma) |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1973 |
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:
74642
diff
changeset
|
1974 |
by auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1975 |
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:
75462
diff
changeset
|
1976 |
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:
74642
diff
changeset
|
1977 |
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:
74642
diff
changeset
|
1978 |
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:
74642
diff
changeset
|
1979 |
finally show ?thesis . |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1980 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1981 |
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:
74642
diff
changeset
|
1982 |
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:
74642
diff
changeset
|
1983 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1984 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1985 |
lemma abs_summable_on_comparison_test: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1986 |
assumes "g abs_summable_on A" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1987 |
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:
74642
diff
changeset
|
1988 |
shows "f abs_summable_on A" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1989 |
proof (rule nonneg_bdd_above_summable_on) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1990 |
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:
74642
diff
changeset
|
1991 |
proof (rule bdd_aboveI2) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1992 |
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:
74642
diff
changeset
|
1993 |
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:
74642
diff
changeset
|
1994 |
using assms F by (intro sum_mono) auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1995 |
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:
74642
diff
changeset
|
1996 |
using F by simp |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
1997 |
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:
76999
diff
changeset
|
1998 |
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:
74642
diff
changeset
|
1999 |
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:
74642
diff
changeset
|
2000 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2001 |
qed auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2002 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2003 |
lemma abs_summable_iff_bdd_above: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2004 |
fixes f :: \<open>'a \<Rightarrow> 'b::real_normed_vector\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2005 |
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:
74642
diff
changeset
|
2006 |
proof (rule iffI) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2007 |
assume \<open>f abs_summable_on A\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2008 |
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:
74642
diff
changeset
|
2009 |
proof (rule bdd_aboveI2) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2010 |
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:
74642
diff
changeset
|
2011 |
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:
74642
diff
changeset
|
2012 |
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:
74642
diff
changeset
|
2013 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2014 |
next |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2015 |
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:
74642
diff
changeset
|
2016 |
then show \<open>f abs_summable_on A\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2017 |
by (simp add: nonneg_bdd_above_summable_on) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2018 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2019 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2020 |
lemma abs_summable_product: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2021 |
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:
74642
diff
changeset
|
2022 |
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:
74642
diff
changeset
|
2023 |
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:
74642
diff
changeset
|
2024 |
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:
74642
diff
changeset
|
2025 |
proof (rule nonneg_bdd_above_summable_on) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2026 |
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:
74642
diff
changeset
|
2027 |
proof (rule bdd_aboveI2) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2028 |
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:
74642
diff
changeset
|
2029 |
then have r1: "finite F" and b4: "F \<subseteq> A" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2030 |
by auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2031 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2032 |
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:
75462
diff
changeset
|
2033 |
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:
74642
diff
changeset
|
2034 |
|
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2035 |
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:
79529
diff
changeset
|
2036 |
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:
74642
diff
changeset
|
2037 |
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:
74642
diff
changeset
|
2038 |
by (simp add: sum_mono) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2039 |
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:
74642
diff
changeset
|
2040 |
by (simp add: sum.distrib) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2041 |
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:
74642
diff
changeset
|
2042 |
by (simp add: \<open>finite F\<close>) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2043 |
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:
74642
diff
changeset
|
2044 |
using F assms |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2045 |
by (intro add_mono infsum_mono2) auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2046 |
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:
74642
diff
changeset
|
2047 |
by simp |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2048 |
qed |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2049 |
qed auto |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2050 |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2051 |
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
|
2052 |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2053 |
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:
76999
diff
changeset
|
2054 |
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
|
2055 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2056 |
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
|
2057 |
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
|
2058 |
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
|
2059 |
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
|
2060 |
assumes \<open>infinite S\<close> |
77359 | 2061 |
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
|
2062 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2063 |
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:
76999
diff
changeset
|
2064 |
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
|
2065 |
fix y :: ennreal assume \<open>y < \<infinity>\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2066 |
then have \<open>y / b < \<infinity>\<close> \<open>y < top\<close> |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2067 |
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
|
2068 |
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
|
2069 |
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
|
2070 |
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
|
2071 |
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
|
2072 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2073 |
have \<open>y < b * card F\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2074 |
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
|
2075 |
also have \<open>\<dots> \<le> b * card Y\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2076 |
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
|
2077 |
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
|
2078 |
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
|
2079 |
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
|
2080 |
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:
74642
diff
changeset
|
2081 |
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
|
2082 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2083 |
ultimately show \<open>\<forall>\<^sub>F x in finite_subsets_at_top S. y < sum f x\<close> |
77359 | 2084 |
unfolding eventually_finite_subsets_at_top by auto |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2085 |
qed auto |
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2086 |
then show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2087 |
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
|
2088 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2089 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2090 |
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
|
2091 |
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
|
2092 |
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
|
2093 |
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
|
2094 |
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
|
2095 |
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
|
2096 |
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
|
2097 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2098 |
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
|
2099 |
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
|
2100 |
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
|
2101 |
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
|
2102 |
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
|
2103 |
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
|
2104 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2105 |
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
|
2106 |
using b by blast |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2107 |
have "0 < e2ennreal b" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2108 |
using b' b |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2109 |
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:
75462
diff
changeset
|
2110 |
hence *: \<open>infsum (e2ennreal \<circ> f) S = \<infinity>\<close> |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2111 |
using assms b' |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2112 |
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:
75462
diff
changeset
|
2113 |
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:
74642
diff
changeset
|
2114 |
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
|
2115 |
also have \<open>\<dots> = enn2ereal \<infinity>\<close> |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2116 |
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
|
2117 |
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
|
2118 |
by simp |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2119 |
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
|
2120 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2121 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2122 |
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
|
2123 |
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
|
2124 |
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
|
2125 |
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
|
2126 |
assumes \<open>infinite S\<close> |
77359 | 2127 |
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
|
2128 |
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
|
2129 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2130 |
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
|
2131 |
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
|
2132 |
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
|
2133 |
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
|
2134 |
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
|
2135 |
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
|
2136 |
proof - |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2137 |
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:
75462
diff
changeset
|
2138 |
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
|
2139 |
also have \<open>\<dots> = \<infinity>\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2140 |
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
|
2141 |
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
|
2142 |
by simp |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2143 |
finally show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2144 |
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
|
2145 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2146 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2147 |
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
|
2148 |
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
|
2149 |
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
|
2150 |
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
|
2151 |
assumes \<open>infinite S\<close> |
77359 | 2152 |
shows "(f has_sum \<infinity>) S" |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2153 |
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
|
2154 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2155 |
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
|
2156 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2157 |
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
|
2158 |
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
|
2159 |
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
|
2160 |
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
|
2161 |
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
|
2162 |
proof - |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2163 |
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:
75462
diff
changeset
|
2164 |
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:
75462
diff
changeset
|
2165 |
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:
75462
diff
changeset
|
2166 |
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
|
2167 |
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:
75462
diff
changeset
|
2168 |
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:
74642
diff
changeset
|
2169 |
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
|
2170 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2171 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2172 |
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
|
2173 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2174 |
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
|
2175 |
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
|
2176 |
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
|
2177 |
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
|
2178 |
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
|
2179 |
proof - |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2180 |
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:
75462
diff
changeset
|
2181 |
by auto |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2182 |
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:
75462
diff
changeset
|
2183 |
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
|
2184 |
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:
74642
diff
changeset
|
2185 |
by (subst nonneg_infsum_complete) (simp_all add: assms) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2186 |
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
|
2187 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2188 |
|
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 |
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
|
2191 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2192 |
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
|
2193 |
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
|
2194 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2195 |
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
|
2196 |
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
|
2197 |
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
|
2198 |
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
|
2199 |
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
|
2200 |
proof - |
77388 | 2201 |
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
|
2202 |
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
|
2203 |
also have "\<dots> = ereal (SUP F\<in>{F. finite F \<and> F \<subseteq> A}. (sum f F))" |
77388 | 2204 |
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
|
2205 |
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
|
2206 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2207 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2208 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2209 |
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
|
2210 |
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
|
2211 |
assumes "f summable_on A" and "\<And>x. x\<in>A \<Longrightarrow> f x \<ge> 0" |
77359 | 2212 |
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
|
2213 |
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
|
2214 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2215 |
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
|
2216 |
fixes f :: \<open>'a \<Rightarrow> real\<close> |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2217 |
assumes \<open>f summable_on A\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2218 |
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:
74642
diff
changeset
|
2219 |
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
|
2220 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2221 |
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
|
2222 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2223 |
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
|
2224 |
fixes f :: \<open>'a \<Rightarrow> complex\<close> |
77359 | 2225 |
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
|
2226 |
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
|
2227 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2228 |
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
|
2229 |
"(\<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
|
2230 |
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
|
2231 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2232 |
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
|
2233 |
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
|
2234 |
|
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2235 |
lemma has_sum_Re: |
77359 | 2236 |
assumes "(f has_sum a) M" |
2237 |
shows "((\<lambda>x. Re (f x)) has_sum Re a) M" |
|
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2238 |
using has_sum_comm_additive[where f=Re] |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2239 |
using assms tendsto_Re by (fastforce simp add: o_def additive_def) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2240 |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2241 |
lemma infsum_Re: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2242 |
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
|
2243 |
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:
75462
diff
changeset
|
2244 |
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
|
2245 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2246 |
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
|
2247 |
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
|
2248 |
shows "(\<lambda>x. Re (f x)) summable_on M" |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2249 |
by (metis assms has_sum_Re summable_on_def) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2250 |
|
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2251 |
lemma has_sum_Im: |
77359 | 2252 |
assumes "(f has_sum a) M" |
2253 |
shows "((\<lambda>x. Im (f x)) has_sum Im a) M" |
|
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2254 |
using has_sum_comm_additive[where f=Im] |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2255 |
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
|
2256 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2257 |
lemma infsum_Im: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2258 |
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
|
2259 |
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:
75462
diff
changeset
|
2260 |
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
|
2261 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2262 |
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
|
2263 |
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
|
2264 |
shows "(\<lambda>x. Im (f x)) summable_on M" |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2265 |
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
|
2266 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2267 |
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
|
2268 |
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
|
2269 |
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
|
2270 |
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
|
2271 |
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
|
2272 |
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
|
2273 |
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
|
2274 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2275 |
have \<open>Im (f x) = 0\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2276 |
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
|
2277 |
moreover have \<open>Re (f x) = 0\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2278 |
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
|
2279 |
ultimately show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2280 |
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
|
2281 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2282 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2283 |
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
|
2284 |
fixes f :: "'a \<Rightarrow> complex" |
77359 | 2285 |
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
|
2286 |
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
|
2287 |
shows "f x = 0" |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2288 |
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
|
2289 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2290 |
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
|
2291 |
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
|
2292 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2293 |
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
|
2294 |
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
|
2295 |
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
|
2296 |
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
|
2297 |
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
|
2298 |
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
|
2299 |
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
|
2300 |
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
|
2301 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2302 |
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:
75462
diff
changeset
|
2303 |
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
|
2304 |
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
|
2305 |
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
|
2306 |
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:
75462
diff
changeset
|
2307 |
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
|
2308 |
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
|
2309 |
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
|
2310 |
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
|
2311 |
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
|
2312 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2313 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2314 |
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
|
2315 |
\<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
|
2316 |
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
|
2317 |
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
|
2318 |
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
|
2319 |
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
|
2320 |
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
|
2321 |
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
|
2322 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2323 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2324 |
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
|
2325 |
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
|
2326 |
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
|
2327 |
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
|
2328 |
shows "infsum f M \<ge> 0" (is "?lhs \<ge> _") |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2329 |
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
|
2330 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2331 |
lemma infsum_cmod: |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2332 |
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
|
2333 |
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
|
2334 |
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
|
2335 |
proof - |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2336 |
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:
74642
diff
changeset
|
2337 |
proof (rule infsum_comm_additive[symmetric, unfolded o_def]) |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2338 |
have "(\<lambda>z. Re (f z)) summable_on M" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2339 |
using assms summable_on_Re by blast |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2340 |
also have "?this \<longleftrightarrow> f abs_summable_on M" |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2341 |
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:
74642
diff
changeset
|
2342 |
finally show \<dots> . |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2343 |
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
|
2344 |
also have \<open>\<dots> = infsum f M\<close> |
76940
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2345 |
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
|
2346 |
finally show ?thesis |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2347 |
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
|
2348 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2349 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2350 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2351 |
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
|
2352 |
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
|
2353 |
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
|
2354 |
proof (rule iffI) |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2355 |
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
|
2356 |
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
|
2357 |
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
|
2358 |
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
|
2359 |
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
|
2360 |
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
|
2361 |
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
|
2362 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2363 |
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
|
2364 |
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
|
2365 |
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
|
2366 |
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
|
2367 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2368 |
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
|
2369 |
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
|
2370 |
|
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2371 |
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:
75462
diff
changeset
|
2372 |
by (simp add: summable_on_add) |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2373 |
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
|
2374 |
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:
75462
diff
changeset
|
2375 |
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:
75462
diff
changeset
|
2376 |
then show \<open>n summable_on A\<close> |
711cef61c0ce
Substantial de-applying and streamlining
paulson <lp15@cam.ac.uk>
parents:
75462
diff
changeset
|
2377 |
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
|
2378 |
next |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2379 |
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
|
2380 |
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
|
2381 |
qed |
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
2382 |
|
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2383 |
lemma summable_countable_complex: |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2384 |
fixes f :: \<open>'a \<Rightarrow> complex\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2385 |
assumes \<open>f summable_on A\<close> |
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
2386 |
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:
74642
diff
changeset
|
2387 |
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:
74642
diff
changeset
|
2388 |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2389 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2390 |
(* TODO: figure all this out *) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2391 |
inductive (in topological_space) convergent_filter :: "'a filter \<Rightarrow> bool" where |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2392 |
"F \<le> nhds x \<Longrightarrow> convergent_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2393 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2394 |
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:
76999
diff
changeset
|
2395 |
by (auto simp: convergent_filter.simps) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2396 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2397 |
lemma (in uniform_space) cauchy_filter_mono: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2398 |
"cauchy_filter F \<Longrightarrow> F' \<le> F \<Longrightarrow> cauchy_filter F'" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2399 |
unfolding cauchy_filter_def by (meson dual_order.trans prod_filter_mono) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2400 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2401 |
lemma (in uniform_space) convergent_filter_cauchy: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2402 |
assumes "convergent_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2403 |
shows "cauchy_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2404 |
using assms cauchy_filter_mono nhds_imp_cauchy_filter[OF order.refl] |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2405 |
by (auto simp: convergent_filter_iff) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2406 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2407 |
lemma (in topological_space) convergent_filter_bot [simp, intro]: "convergent_filter bot" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2408 |
by (simp add: convergent_filter_iff) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2409 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2410 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2411 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2412 |
class complete_uniform_space = uniform_space + |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2413 |
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:
76999
diff
changeset
|
2414 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2415 |
lemma (in complete_uniform_space) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2416 |
cauchy_filter_convergent: "cauchy_filter (F :: 'a filter) \<Longrightarrow> convergent_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2417 |
using cauchy_filter_convergent'[of F] by (cases "F = bot") auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2418 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2419 |
lemma (in complete_uniform_space) convergent_filter_iff_cauchy: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2420 |
"convergent_filter F \<longleftrightarrow> cauchy_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2421 |
using convergent_filter_cauchy cauchy_filter_convergent by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2422 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2423 |
definition countably_generated_filter :: "'a filter \<Rightarrow> bool" where |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2424 |
"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:
76999
diff
changeset
|
2425 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2426 |
lemma countably_generated_filter_has_antimono_basis: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2427 |
assumes "countably_generated_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2428 |
obtains B :: "nat \<Rightarrow> 'a set" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2429 |
where "antimono B" and "F = (INF n. principal (B n))" and |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2430 |
"\<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:
76999
diff
changeset
|
2431 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2432 |
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:
76999
diff
changeset
|
2433 |
unfolding countably_generated_filter_def by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2434 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2435 |
define B' where "B' = (\<lambda>n. \<Inter>k\<le>n. B k)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2436 |
have "antimono B'" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2437 |
unfolding decseq_def B'_def by force |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2438 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2439 |
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:
76999
diff
changeset
|
2440 |
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:
76999
diff
changeset
|
2441 |
also have "\<dots> = (INF (n::nat). principal (B n))" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2442 |
apply (intro antisym) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2443 |
apply (meson INF_lower INF_mono atMost_iff order_refl) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2444 |
apply (meson INF_greatest INF_lower UNIV_I) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2445 |
done |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2446 |
also have "\<dots> = F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2447 |
by (simp add: B) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2448 |
finally have F: "F = (INF n. principal (B' n))" .. |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2449 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2450 |
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:
76999
diff
changeset
|
2451 |
unfolding F using \<open>antimono B'\<close> |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2452 |
apply (subst eventually_INF_base) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2453 |
apply (auto simp: decseq_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2454 |
by (meson nat_le_linear) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2455 |
ultimately show ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2456 |
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:
76999
diff
changeset
|
2457 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2458 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2459 |
lemma (in uniform_space) cauchy_filter_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2460 |
"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:
76999
diff
changeset
|
2461 |
unfolding cauchy_filter_def le_filter_def |
79529
cb933e165dc3
tuned proof: avoid z3 to make it work on arm64-linux;
wenzelm
parents:
77388
diff
changeset
|
2462 |
apply (auto simp: eventually_prod_same) |
cb933e165dc3
tuned proof: avoid z3 to make it work on arm64-linux;
wenzelm
parents:
77388
diff
changeset
|
2463 |
apply (metis (full_types) eventually_mono mem_Collect_eq) |
cb933e165dc3
tuned proof: avoid z3 to make it work on arm64-linux;
wenzelm
parents:
77388
diff
changeset
|
2464 |
apply blast |
cb933e165dc3
tuned proof: avoid z3 to make it work on arm64-linux;
wenzelm
parents:
77388
diff
changeset
|
2465 |
done |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2466 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2467 |
lemma (in uniform_space) controlled_sequences_convergent_imp_complete_aux_sequence: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2468 |
fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2469 |
fixes F :: "'a filter" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2470 |
assumes "cauchy_filter F" "F \<noteq> bot" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2471 |
assumes "\<And>n. eventually (\<lambda>z. z \<in> U n) uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2472 |
obtains g G where |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2473 |
"antimono G" "\<And>n. g n \<in> G n" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2474 |
"\<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:
76999
diff
changeset
|
2475 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2476 |
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:
76999
diff
changeset
|
2477 |
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:
76999
diff
changeset
|
2478 |
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:
76999
diff
changeset
|
2479 |
by metis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2480 |
define G' where "G' = (\<lambda>n. \<Inter>k\<le>n. G k)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2481 |
have 1: "eventually (\<lambda>x. x \<in> G' n) F" for n |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2482 |
using G by (auto simp: G'_def intro: eventually_ball_finite) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2483 |
have 2: "G' n \<times> G' n \<subseteq> U n" for n |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2484 |
using G unfolding G'_def by fast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2485 |
have 3: "antimono G'" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2486 |
unfolding G'_def decseq_def by force |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2487 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2488 |
have "\<exists>g. g \<in> G' n" for n |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2489 |
using 1 assms(2) eventually_happens' by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2490 |
then obtain g where g: "\<And>n. g n \<in> G' n" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2491 |
by metis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2492 |
from g 1 2 3 that[of G' g] show ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2493 |
by metis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2494 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2495 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2496 |
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:
76999
diff
changeset
|
2497 |
"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:
76999
diff
changeset
|
2498 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2499 |
lemma lift_filter_top [simp]: "lift_filter g top = g UNIV" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2500 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2501 |
have "{X. \<forall>x::'b. x \<in> X} = {UNIV}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2502 |
by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2503 |
thus ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2504 |
by (simp add: lift_filter_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2505 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2506 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2507 |
lemma eventually_lift_filter_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2508 |
assumes "mono g" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2509 |
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:
76999
diff
changeset
|
2510 |
unfolding lift_filter_def |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2511 |
proof (subst eventually_INF_base, goal_cases) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2512 |
case 1 |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2513 |
thus ?case by (auto intro: exI[of _ UNIV]) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2514 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2515 |
case (2 X Y) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2516 |
thus ?case |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2517 |
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:
76999
diff
changeset
|
2518 |
qed auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2519 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2520 |
lemma lift_filter_le: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2521 |
assumes "eventually (\<lambda>x. x \<in> X) F" "g X \<le> F'" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2522 |
shows "lift_filter g F \<le> F'" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2523 |
unfolding lift_filter_def |
77388 | 2524 |
by (metis INF_lower2 assms mem_Collect_eq) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2525 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2526 |
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:
76999
diff
changeset
|
2527 |
"lift_filter' f F = lift_filter (principal \<circ> f) F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2528 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2529 |
lemma lift_filter'_top [simp]: "lift_filter' g top = principal (g UNIV)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2530 |
by (simp add: lift_filter'_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2531 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2532 |
lemma eventually_lift_filter'_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2533 |
assumes "mono g" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2534 |
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:
76999
diff
changeset
|
2535 |
unfolding lift_filter'_def using assms |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2536 |
by (subst eventually_lift_filter_iff) (auto simp: mono_def eventually_principal) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2537 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2538 |
lemma lift_filter'_le: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2539 |
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:
76999
diff
changeset
|
2540 |
shows "lift_filter' g F \<le> F'" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2541 |
unfolding lift_filter'_def using assms |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2542 |
by (intro lift_filter_le[where X = X]) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2543 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2544 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2545 |
lemma (in uniform_space) comp_uniformity_le_uniformity: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2546 |
"lift_filter' (\<lambda>X. X O X) uniformity \<le> uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2547 |
unfolding le_filter_def |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2548 |
proof safe |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2549 |
fix P assume P: "eventually P uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2550 |
have [simp]: "mono (\<lambda>X::('a \<times> 'a) set. X O X)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2551 |
by (intro monoI) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2552 |
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:
76999
diff
changeset
|
2553 |
using uniformity_transE by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2554 |
show "eventually P (lift_filter' (\<lambda>X. X O X) uniformity)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2555 |
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:
76999
diff
changeset
|
2556 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2557 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2558 |
lemma (in uniform_space) comp_mem_uniformity_sets: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2559 |
assumes "eventually (\<lambda>z. z \<in> X) uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2560 |
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:
76999
diff
changeset
|
2561 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2562 |
have [simp]: "mono (\<lambda>X::('a \<times> 'a) set. X O X)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2563 |
by (intro monoI) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2564 |
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:
76999
diff
changeset
|
2565 |
using assms comp_uniformity_le_uniformity using filter_leD by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2566 |
thus ?thesis using that |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2567 |
by (auto simp: eventually_lift_filter'_iff) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2568 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2569 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2570 |
lemma (in uniform_space) le_nhds_of_cauchy_adhp_aux: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2571 |
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:
76999
diff
changeset
|
2572 |
shows "F \<le> nhds x" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2573 |
unfolding le_filter_def |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2574 |
proof safe |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2575 |
fix P assume "eventually P (nhds x)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2576 |
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:
76999
diff
changeset
|
2577 |
by (simp add: eventually_nhds_uniformity case_prod_unfold) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2578 |
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:
76999
diff
changeset
|
2579 |
using comp_mem_uniformity_sets by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2580 |
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:
76999
diff
changeset
|
2581 |
using assms[OF Y(1)] by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2582 |
have *: "P x" if "x \<in> X" for x |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2583 |
using Y(2) Xy(2-4) that unfolding relcomp_unfold by force |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2584 |
show "eventually P F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2585 |
by (rule eventually_mono[OF Xy(1)]) (use * in auto) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2586 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2587 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2588 |
lemma (in uniform_space) eventually_uniformity_imp_nhds: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2589 |
assumes "eventually P uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2590 |
shows "eventually (\<lambda>y. P (x, y)) (nhds x)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2591 |
using assms unfolding eventually_nhds_uniformity by (elim eventually_mono) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2592 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2593 |
lemma (in uniform_space) controlled_sequences_convergent_imp_complete_aux: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2594 |
fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2595 |
assumes gen: "countably_generated_filter (uniformity :: ('a \<times> 'a) filter)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2596 |
assumes U: "\<And>n. eventually (\<lambda>z. z \<in> U n) uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2597 |
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:
76999
diff
changeset
|
2598 |
assumes "cauchy_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2599 |
shows "convergent_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2600 |
proof (cases "F = bot") |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2601 |
case False |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2602 |
note F = \<open>cauchy_filter F\<close> \<open>F \<noteq> bot\<close> |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2603 |
from gen obtain B :: "nat \<Rightarrow> ('a \<times> 'a) set" where B: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2604 |
"antimono B" "uniformity = (INF n. principal (B n))" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2605 |
"\<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:
76999
diff
changeset
|
2606 |
using countably_generated_filter_has_antimono_basis by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2607 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2608 |
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:
76999
diff
changeset
|
2609 |
by (subst B(3)) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2610 |
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:
76999
diff
changeset
|
2611 |
using U by (auto intro: eventually_conj) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2612 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2613 |
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:
76999
diff
changeset
|
2614 |
"\<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:
76999
diff
changeset
|
2615 |
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:
76999
diff
changeset
|
2616 |
by metis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2617 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2618 |
have "convergent g" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2619 |
proof (rule conv) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2620 |
fix N m n :: nat |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2621 |
assume mn: "N \<le> m" "N \<le> n" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2622 |
have "(g m, g n) \<in> G m \<times> G n" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2623 |
using gG by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2624 |
also from mn have "\<dots> \<subseteq> G N \<times> G N" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2625 |
by (intro Sigma_mono gG antimonoD[OF gG(1)]) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2626 |
also have "\<dots> \<subseteq> U N" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2627 |
using gG by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2628 |
finally show "(g m, g n) \<in> U N" . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2629 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2630 |
then obtain L where G: "g \<longlonglongrightarrow> L" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2631 |
unfolding convergent_def by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2632 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2633 |
have "F \<le> nhds L" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2634 |
proof (rule le_nhds_of_cauchy_adhp_aux) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2635 |
fix P :: "'a \<times> 'a \<Rightarrow> bool" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2636 |
assume P: "eventually P uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2637 |
hence "eventually (\<lambda>n. \<forall>x\<in>B n. P x) sequentially" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2638 |
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:
76999
diff
changeset
|
2639 |
moreover have "eventually (\<lambda>n. P (L, g n)) sequentially" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2640 |
using P eventually_compose_filterlim eventually_uniformity_imp_nhds G by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2641 |
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:
76999
diff
changeset
|
2642 |
by eventually_elim auto |
77388 | 2643 |
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:
76999
diff
changeset
|
2644 |
unfolding eventually_at_top_linorder by blast |
77388 | 2645 |
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)" |
2646 |
using gG by blast+ |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2647 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2648 |
thus "convergent_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2649 |
by (auto simp: convergent_filter_iff) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2650 |
qed auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2651 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2652 |
theorem (in uniform_space) controlled_sequences_convergent_imp_complete: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2653 |
fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2654 |
assumes gen: "countably_generated_filter (uniformity :: ('a \<times> 'a) filter)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2655 |
assumes U: "\<And>n. eventually (\<lambda>z. z \<in> U n) uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2656 |
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:
76999
diff
changeset
|
2657 |
shows "class.complete_uniform_space open uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2658 |
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:
76999
diff
changeset
|
2659 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2660 |
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:
76999
diff
changeset
|
2661 |
proof (intro antisym) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2662 |
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:
76999
diff
changeset
|
2663 |
by (auto simp: le_filter_def eventually_filtermap eventually_prod_filter) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2664 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2665 |
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:
76999
diff
changeset
|
2666 |
unfolding le_filter_def |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2667 |
proof safe |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2668 |
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:
76999
diff
changeset
|
2669 |
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:
76999
diff
changeset
|
2670 |
by (auto simp: eventually_filtermap eventually_prod_filter) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2671 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2672 |
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:
76999
diff
changeset
|
2673 |
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:
76999
diff
changeset
|
2674 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2675 |
from *(1) have "\<forall>\<^sub>F x in F. Pf' (f x)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2676 |
by eventually_elim (auto simp: Pf'_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2677 |
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:
76999
diff
changeset
|
2678 |
by eventually_elim (auto simp: Pg'_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2679 |
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:
76999
diff
changeset
|
2680 |
using *(3) by (auto simp: Pf'_def Pg'_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2681 |
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:
76999
diff
changeset
|
2682 |
unfolding eventually_prod_filter eventually_filtermap |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2683 |
by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2684 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2685 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2686 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2687 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2688 |
lemma (in uniform_space) Cauchy_seq_iff_tendsto: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2689 |
"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:
76999
diff
changeset
|
2690 |
unfolding Cauchy_uniform cauchy_filter_def filterlim_def filtermap_prod_filter .. |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2691 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2692 |
theorem (in uniform_space) controlled_seq_imp_Cauchy_seq: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2693 |
fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2694 |
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:
76999
diff
changeset
|
2695 |
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:
76999
diff
changeset
|
2696 |
shows "Cauchy f" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2697 |
unfolding Cauchy_seq_iff_tendsto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2698 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2699 |
show "filterlim (map_prod f f) uniformity (sequentially \<times>\<^sub>F sequentially)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2700 |
unfolding filterlim_def le_filter_def |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2701 |
proof safe |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2702 |
fix P :: "'a \<times> 'a \<Rightarrow> bool" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2703 |
assume P: "eventually P uniformity" |
77388 | 2704 |
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:
76999
diff
changeset
|
2705 |
by blast |
77388 | 2706 |
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:
76999
diff
changeset
|
2707 |
unfolding eventually_filtermap eventually_prod_sequentially |
77388 | 2708 |
by (metis controlled map_prod_simp) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2709 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2710 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2711 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2712 |
lemma (in uniform_space) Cauchy_seq_convergent_imp_complete_aux: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2713 |
fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2714 |
assumes gen: "countably_generated_filter (uniformity :: ('a \<times> 'a) filter)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2715 |
assumes conv: "\<And>(u :: nat \<Rightarrow> 'a). Cauchy u \<Longrightarrow> convergent u" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2716 |
assumes "cauchy_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2717 |
shows "convergent_filter F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2718 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2719 |
from gen obtain B :: "nat \<Rightarrow> ('a \<times> 'a) set" where B: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2720 |
"antimono B" "uniformity = (INF n. principal (B n))" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2721 |
"\<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:
76999
diff
changeset
|
2722 |
using countably_generated_filter_has_antimono_basis by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2723 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2724 |
show ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2725 |
proof (rule controlled_sequences_convergent_imp_complete_aux[where U = B]) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2726 |
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:
76999
diff
changeset
|
2727 |
unfolding B(3) by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2728 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2729 |
fix f :: "nat \<Rightarrow> 'a" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2730 |
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:
76999
diff
changeset
|
2731 |
have "Cauchy f" using f B |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2732 |
by (intro controlled_seq_imp_Cauchy_seq[where U = B]) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2733 |
with conv show "convergent f" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2734 |
by simp |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2735 |
qed fact+ |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2736 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2737 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2738 |
theorem (in uniform_space) Cauchy_seq_convergent_imp_complete: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2739 |
fixes U :: "nat \<Rightarrow> ('a \<times> 'a) set" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2740 |
assumes gen: "countably_generated_filter (uniformity :: ('a \<times> 'a) filter)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2741 |
assumes conv: "\<And>(u :: nat \<Rightarrow> 'a). Cauchy u \<Longrightarrow> convergent u" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2742 |
shows "class.complete_uniform_space open uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2743 |
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:
76999
diff
changeset
|
2744 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2745 |
lemma (in metric_space) countably_generated_uniformity: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2746 |
"countably_generated_filter uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2747 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2748 |
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:
76999
diff
changeset
|
2749 |
(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:
76999
diff
changeset
|
2750 |
unfolding uniformity_dist |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2751 |
proof (intro antisym) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2752 |
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:
76999
diff
changeset
|
2753 |
by (simp add: image_image) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2754 |
also have "\<dots> \<ge> ?F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2755 |
by (intro INF_superset_mono) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2756 |
finally show "?F \<le> ?G" . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2757 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2758 |
show "?G \<le> ?F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2759 |
unfolding le_filter_def |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2760 |
proof safe |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2761 |
fix P assume "eventually P ?F" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2762 |
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:
76999
diff
changeset
|
2763 |
proof (subst (asm) eventually_INF_base, goal_cases) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2764 |
case (2 \<epsilon>1 \<epsilon>2) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2765 |
thus ?case |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2766 |
by (intro bexI[of _ "min \<epsilon>1 \<epsilon>2"]) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2767 |
qed auto |
77388 | 2768 |
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:
76999
diff
changeset
|
2769 |
using nat_approx_posE by blast |
77388 | 2770 |
then have "eventually P (principal {(x, y). dist x y < 1 / real (Suc n)})" |
2771 |
using \<epsilon>(2) by (auto simp: eventually_principal) |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2772 |
thus "eventually P ?G" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2773 |
by (intro eventually_INF1) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2774 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2775 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2776 |
thus "countably_generated_filter uniformity" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2777 |
unfolding countably_generated_filter_def uniformity_dist by fast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2778 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2779 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2780 |
subclass (in complete_space) complete_uniform_space |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2781 |
proof (rule Cauchy_seq_convergent_imp_complete) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2782 |
show "convergent f" if "Cauchy f" for f |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2783 |
using Cauchy_convergent that by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2784 |
qed (fact countably_generated_uniformity) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2785 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2786 |
lemma (in complete_uniform_space) complete_UNIV_cuspace [intro]: "complete UNIV" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2787 |
unfolding complete_uniform using cauchy_filter_convergent |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2788 |
by (auto simp: convergent_filter.simps) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2789 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2790 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2791 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2792 |
lemma norm_infsum_le: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2793 |
assumes "(f has_sum S) X" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2794 |
assumes "(g has_sum T) X" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2795 |
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:
76999
diff
changeset
|
2796 |
shows "norm S \<le> T" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2797 |
proof (rule tendsto_le) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2798 |
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:
76999
diff
changeset
|
2799 |
using assms(1) unfolding has_sum_def by (intro tendsto_norm) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2800 |
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:
76999
diff
changeset
|
2801 |
using assms(2) unfolding has_sum_def . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2802 |
show "\<forall>\<^sub>F x in finite_subsets_at_top X. norm (sum f x) \<le> (\<Sum>x\<in>x. g x)" |
77388 | 2803 |
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:
76999
diff
changeset
|
2804 |
qed auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2805 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2806 |
(* |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2807 |
lemma summable_on_Sigma: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2808 |
fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2809 |
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:
76999
diff
changeset
|
2810 |
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:
76999
diff
changeset
|
2811 |
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:
76999
diff
changeset
|
2812 |
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:
76999
diff
changeset
|
2813 |
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:
76999
diff
changeset
|
2814 |
*) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2815 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2816 |
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:
76999
diff
changeset
|
2817 |
by (auto simp: summable_on_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2818 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2819 |
lemma has_sum_reindex_bij_betw: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2820 |
assumes "bij_betw g A B" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2821 |
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:
76999
diff
changeset
|
2822 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2823 |
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:
76999
diff
changeset
|
2824 |
by (subst has_sum_reindex) (use assms in \<open>auto dest: bij_betw_imp_inj_on simp: o_def\<close>) |
77388 | 2825 |
then show ?thesis |
2826 |
using assms bij_betw_imp_surj_on by blast |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2827 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2828 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2829 |
lemma has_sum_reindex_bij_witness: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2830 |
assumes "\<And>a. a \<in> S \<Longrightarrow> i (j a) = a" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2831 |
assumes "\<And>a. a \<in> S \<Longrightarrow> j a \<in> T" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2832 |
assumes "\<And>b. b \<in> T \<Longrightarrow> j (i b) = b" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2833 |
assumes "\<And>b. b \<in> T \<Longrightarrow> i b \<in> S" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2834 |
assumes "\<And>a. a \<in> S \<Longrightarrow> h (j a) = g a" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2835 |
assumes "s = s'" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2836 |
shows "(g has_sum s) S = (h has_sum s') T" |
77388 | 2837 |
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:
76999
diff
changeset
|
2838 |
|
82349
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2839 |
lemma summable_on_reindex_bij_witness: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2840 |
assumes "\<And>a. a \<in> S \<Longrightarrow> i (j a) = a" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2841 |
assumes "\<And>a. a \<in> S \<Longrightarrow> j a \<in> T" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2842 |
assumes "\<And>b. b \<in> T \<Longrightarrow> j (i b) = b" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2843 |
assumes "\<And>b. b \<in> T \<Longrightarrow> i b \<in> S" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2844 |
assumes "\<And>a. a \<in> S \<Longrightarrow> h (j a) = g a" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2845 |
shows "g summable_on S \<longleftrightarrow> h summable_on T" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2846 |
using has_sum_reindex_bij_witness[of S i j T h g, OF assms] |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2847 |
by (simp add: summable_on_def) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2848 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2849 |
lemma has_sum_homomorphism: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2850 |
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:
76999
diff
changeset
|
2851 |
shows "((\<lambda>x. h (f x)) has_sum (h S)) A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2852 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2853 |
have "sum (h \<circ> f) X = h (sum f X)" for X |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2854 |
by (induction X rule: infinite_finite_induct) (simp_all add: assms) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2855 |
hence sum_h: "sum (h \<circ> f) = h \<circ> sum f" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2856 |
by (intro ext) auto |
77388 | 2857 |
then have "((h \<circ> f) has_sum h S) A" |
2858 |
using assms |
|
2859 |
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:
76999
diff
changeset
|
2860 |
thus ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2861 |
by (simp add: o_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2862 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2863 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2864 |
lemma summable_on_homomorphism: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2865 |
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:
76999
diff
changeset
|
2866 |
shows "(\<lambda>x. h (f x)) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2867 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2868 |
from assms(1) obtain S where "(f has_sum S) A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2869 |
by (auto simp: summable_on_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2870 |
hence "((\<lambda>x. h (f x)) has_sum h S) A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2871 |
by (rule has_sum_homomorphism) (use assms in auto) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2872 |
thus ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2873 |
by (auto simp: summable_on_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2874 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2875 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2876 |
lemma infsum_homomorphism_strong: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2877 |
fixes h :: "'a :: {t2_space, topological_comm_monoid_add} \<Rightarrow> |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2878 |
'b :: {t2_space, topological_comm_monoid_add}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2879 |
assumes "(\<lambda>x. h (f x)) summable_on A \<longleftrightarrow> f summable_on A" |
77388 | 2880 |
assumes "h 0 = 0" |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2881 |
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:
76999
diff
changeset
|
2882 |
shows "infsum (\<lambda>x. h (f x)) A = h (infsum f A)" |
77388 | 2883 |
by (metis assms has_sum_infsum infsumI infsum_not_exists) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2884 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2885 |
lemma has_sum_bounded_linear: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2886 |
assumes "bounded_linear h" and "(f has_sum S) A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2887 |
shows "((\<lambda>x. h (f x)) has_sum h S) A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2888 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2889 |
interpret bounded_linear h by fact |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2890 |
from assms(2) show ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2891 |
by (rule has_sum_homomorphism) (auto simp: add intro!: continuous_on) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2892 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2893 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2894 |
lemma summable_on_bounded_linear: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2895 |
assumes "bounded_linear h" and "f summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2896 |
shows "(\<lambda>x. h (f x)) summable_on A" |
77388 | 2897 |
by (metis assms has_sum_bounded_linear summable_on_def) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2898 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2899 |
lemma summable_on_bounded_linear_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2900 |
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:
76999
diff
changeset
|
2901 |
shows "(\<lambda>x. h (f x)) summable_on A \<longleftrightarrow> f summable_on A" |
77388 | 2902 |
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:
76999
diff
changeset
|
2903 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2904 |
lemma infsum_bounded_linear_strong: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2905 |
fixes h :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_vector" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2906 |
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:
76999
diff
changeset
|
2907 |
assumes "bounded_linear h" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2908 |
shows "infsum (\<lambda>x. h (f x)) A = h (infsum f A)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2909 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2910 |
interpret bounded_linear h by fact |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2911 |
show ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2912 |
by (rule infsum_homomorphism_strong) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2913 |
(insert assms, auto intro: add continuous_on has_sum_bounded_linear) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2914 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2915 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2916 |
lemma infsum_bounded_linear_strong': |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2917 |
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:
76999
diff
changeset
|
2918 |
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:
76999
diff
changeset
|
2919 |
assumes "bounded_linear (mult c)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2920 |
assumes [simp]: "\<And>x. mult 0 x = 0" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2921 |
shows "infsum (\<lambda>x. mult c (f x)) A = mult c (infsum f A)" |
77388 | 2922 |
by (metis assms infsum_0 infsum_bounded_linear_strong) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2923 |
|
82349
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2924 |
lemma has_sum_scaleR: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2925 |
fixes f :: "'a \<Rightarrow> 'b :: real_normed_vector" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2926 |
assumes "(f has_sum S) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2927 |
shows "((\<lambda>x. c *\<^sub>R f x) has_sum (c *\<^sub>R S)) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2928 |
using has_sum_bounded_linear[OF bounded_linear_scaleR_right[of c], of f A S] assms by simp |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2929 |
|
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2930 |
lemma has_sum_scaleR_iff: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2931 |
fixes f :: "'a \<Rightarrow> 'b :: real_normed_vector" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2932 |
assumes "c \<noteq> 0" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2933 |
shows "((\<lambda>x. c *\<^sub>R f x) has_sum S) A \<longleftrightarrow> (f has_sum (S /\<^sub>R c)) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2934 |
using has_sum_scaleR[of f A "S /\<^sub>R c" c] has_sum_scaleR[of "\<lambda>x. c *\<^sub>R f x" A S "inverse c"] assms |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2935 |
by auto |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2936 |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2937 |
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:
76999
diff
changeset
|
2938 |
by (erule has_sum_homomorphism) (auto intro!: continuous_intros) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2939 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2940 |
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:
76999
diff
changeset
|
2941 |
by (erule has_sum_homomorphism) (auto intro!: continuous_intros) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2942 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2943 |
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:
76999
diff
changeset
|
2944 |
by (erule summable_on_homomorphism) (auto intro!: continuous_intros) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2945 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2946 |
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:
76999
diff
changeset
|
2947 |
by (erule summable_on_homomorphism) (auto intro!: continuous_intros) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2948 |
|
82349
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2949 |
lemma summable_on_of_real: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2950 |
"f summable_on A \<Longrightarrow> (\<lambda>x. of_real (f x) :: 'a :: real_normed_algebra_1) summable_on A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2951 |
using summable_on_bounded_linear[of "of_real :: real \<Rightarrow> 'a", OF bounded_linear_of_real, of f A] |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2952 |
by simp |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2953 |
|
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2954 |
lemma has_sum_of_real_iff: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2955 |
"((\<lambda>x. of_real (f x) :: 'a :: real_normed_div_algebra) has_sum (of_real c)) A \<longleftrightarrow> |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2956 |
(f has_sum c) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2957 |
proof - |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2958 |
have "((\<lambda>x. of_real (f x) :: 'a) has_sum (of_real c)) A \<longleftrightarrow> |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2959 |
(sum (\<lambda>x. of_real (f x) :: 'a) \<longlongrightarrow> of_real c) (finite_subsets_at_top A)" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2960 |
by (simp add: has_sum_def) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2961 |
also have "sum (\<lambda>x. of_real (f x) :: 'a) = (\<lambda>X. of_real (sum f X))" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2962 |
by simp |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2963 |
also have "((\<lambda>X. of_real (sum f X) :: 'a) \<longlongrightarrow> of_real c) (finite_subsets_at_top A) \<longleftrightarrow> |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2964 |
(f has_sum c) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2965 |
unfolding has_sum_def tendsto_of_real_iff .. |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2966 |
finally show ?thesis . |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2967 |
qed |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2968 |
|
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2969 |
lemma has_sum_of_real: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2970 |
"(f has_sum S) A \<Longrightarrow> ((\<lambda>x. of_real (f x) :: 'a :: real_normed_algebra_1) has_sum of_real S) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2971 |
using has_sum_bounded_linear[of "of_real :: real \<Rightarrow> 'a", OF bounded_linear_of_real, of f A S] |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2972 |
by simp |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
2973 |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2974 |
lemma summable_on_discrete_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2975 |
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:
76999
diff
changeset
|
2976 |
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:
76999
diff
changeset
|
2977 |
proof |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2978 |
assume *: "finite {x\<in>A. f x \<noteq> 0}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2979 |
hence "f summable_on {x\<in>A. f x \<noteq> 0}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2980 |
by (rule summable_on_finite) |
77388 | 2981 |
then show "f summable_on A" |
2982 |
by (smt (verit) DiffE mem_Collect_eq summable_on_cong_neutral) |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2983 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2984 |
assume "f summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2985 |
then obtain S where "(f has_sum S) A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2986 |
by (auto simp: summable_on_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2987 |
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:
76999
diff
changeset
|
2988 |
unfolding has_sum_def tendsto_discrete . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2989 |
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:
76999
diff
changeset
|
2990 |
unfolding eventually_finite_subsets_at_top by metis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2991 |
have "{x\<in>A. f x \<noteq> 0} \<subseteq> X" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2992 |
proof |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2993 |
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:
76999
diff
changeset
|
2994 |
show "x \<in> X" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2995 |
proof (rule ccontr) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2996 |
assume [simp]: "x \<notin> X" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
2997 |
have "sum f (insert x X) = S" |
77388 | 2998 |
using X x by (intro X) auto |
2999 |
then have "f x = 0" |
|
3000 |
using X by auto |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3001 |
with x show False |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3002 |
by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3003 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3004 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3005 |
thus "finite {x\<in>A. f x \<noteq> 0}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3006 |
using X(1) finite_subset by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3007 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3008 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3009 |
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:
76999
diff
changeset
|
3010 |
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:
76999
diff
changeset
|
3011 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3012 |
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:
76999
diff
changeset
|
3013 |
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:
76999
diff
changeset
|
3014 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3015 |
lemma summable_on_UNIV_nonneg_real_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3016 |
assumes "\<And>n. f n \<ge> (0 :: real)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3017 |
shows "f summable_on UNIV \<longleftrightarrow> summable f" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3018 |
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:
76999
diff
changeset
|
3019 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3020 |
lemma summable_on_imp_bounded_partial_sums: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3021 |
fixes f :: "_ \<Rightarrow> 'a :: {topological_comm_monoid_add, linorder_topology}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3022 |
assumes f: "f summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3023 |
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:
76999
diff
changeset
|
3024 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3025 |
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:
76999
diff
changeset
|
3026 |
unfolding summable_on_def has_sum_def by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3027 |
show ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3028 |
proof (cases "\<exists>C. C > S") |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3029 |
case True |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3030 |
then obtain C where C: "C > S" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3031 |
by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3032 |
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:
76999
diff
changeset
|
3033 |
using S C by (rule order_tendstoD(2)) |
77388 | 3034 |
thus ?thesis |
3035 |
by (meson eventually_mono nless_le) |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3036 |
next |
77388 | 3037 |
case False thus ?thesis |
3038 |
by (meson not_eventuallyD not_le_imp_less) |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3039 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3040 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3041 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3042 |
lemma has_sum_mono': |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3043 |
fixes S S' :: "'a :: {linorder_topology, ordered_comm_monoid_add, topological_comm_monoid_add}" |
77388 | 3044 |
assumes f: "(f has_sum S) A" "(f has_sum S') B" |
3045 |
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:
76999
diff
changeset
|
3046 |
shows "S \<le> S'" |
77388 | 3047 |
using AB has_sum_mono_neutral[OF f] by fastforce |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3048 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3049 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3050 |
context |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3051 |
assumes "SORT_CONSTRAINT('a :: {topological_comm_monoid_add, order_topology, |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3052 |
ordered_comm_monoid_add, conditionally_complete_linorder})" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3053 |
begin |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3054 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3055 |
text \<open> |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3056 |
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:
76999
diff
changeset
|
3057 |
is simply the supremum of the partial sums. |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3058 |
\<close> |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3059 |
lemma nonneg_bounded_partial_sums_imp_has_sum_SUP: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3060 |
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:
76999
diff
changeset
|
3061 |
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:
76999
diff
changeset
|
3062 |
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:
76999
diff
changeset
|
3063 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3064 |
from bound obtain X0 |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3065 |
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:
76999
diff
changeset
|
3066 |
by (force simp: eventually_finite_subsets_at_top) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3067 |
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:
76999
diff
changeset
|
3068 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3069 |
have "sum f X \<le> sum f (X \<union> X0)" |
77388 | 3070 |
using that X0 assms(1) by (intro sum_mono2) auto |
3071 |
also have "\<dots> \<le> C" |
|
3072 |
by (simp add: X0 that) |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3073 |
finally show ?thesis . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3074 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3075 |
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:
76999
diff
changeset
|
3076 |
by (auto simp: bdd_above_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3077 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3078 |
show ?thesis unfolding has_sum_def |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3079 |
proof (rule increasing_tendsto) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3080 |
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:
76999
diff
changeset
|
3081 |
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:
76999
diff
changeset
|
3082 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3083 |
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:
76999
diff
changeset
|
3084 |
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:
76999
diff
changeset
|
3085 |
by (subst (asm) less_cSUP_iff[OF _ bdd]) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3086 |
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:
76999
diff
changeset
|
3087 |
by (auto simp: eventually_finite_subsets_at_top) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3088 |
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:
76999
diff
changeset
|
3089 |
proof eventually_elim |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3090 |
case (elim X') |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3091 |
note \<open>y < sum f X\<close> |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3092 |
also have "sum f X \<le> sum f X'" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3093 |
using nonneg elim by (intro sum_mono2) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3094 |
finally show ?case . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3095 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3096 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3097 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3098 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3099 |
lemma nonneg_bounded_partial_sums_imp_summable_on: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3100 |
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:
76999
diff
changeset
|
3101 |
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:
76999
diff
changeset
|
3102 |
shows "f summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3103 |
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:
76999
diff
changeset
|
3104 |
|
74475
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff
changeset
|
3105 |
end |
74791
227915e07891
more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents:
74642
diff
changeset
|
3106 |
|
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3107 |
context |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3108 |
assumes "SORT_CONSTRAINT('a :: {topological_comm_monoid_add, linorder_topology, |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3109 |
ordered_comm_monoid_add, conditionally_complete_linorder})" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3110 |
begin |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3111 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3112 |
lemma summable_on_comparison_test: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3113 |
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:
76999
diff
changeset
|
3114 |
shows "g summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3115 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3116 |
obtain C where C: "\<forall>\<^sub>F X in finite_subsets_at_top A. sum f X \<le> C" |
77388 | 3117 |
using assms(1) summable_on_imp_bounded_partial_sums by blast |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3118 |
show ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3119 |
proof (rule nonneg_bounded_partial_sums_imp_summable_on) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3120 |
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:
76999
diff
changeset
|
3121 |
using C assms |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3122 |
unfolding eventually_finite_subsets_at_top |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3123 |
by (smt (verit, ccfv_SIG) order_trans subsetD sum_mono) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3124 |
qed (use assms in auto) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3125 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3126 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3127 |
end |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3128 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3129 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3130 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3131 |
lemma summable_on_subset: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3132 |
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:
76999
diff
changeset
|
3133 |
assumes "f summable_on A" "B \<subseteq> A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3134 |
shows "f summable_on B" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3135 |
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:
76999
diff
changeset
|
3136 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3137 |
lemma summable_on_union: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3138 |
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:
76999
diff
changeset
|
3139 |
assumes "f summable_on A" "f summable_on B" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3140 |
shows "f summable_on (A \<union> B)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3141 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3142 |
have "f summable_on (A \<union> (B - A))" |
77388 | 3143 |
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:
76999
diff
changeset
|
3144 |
also have "A \<union> (B - A) = A \<union> B" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3145 |
by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3146 |
finally show ?thesis . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3147 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3148 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3149 |
lemma summable_on_insert_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3150 |
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:
76999
diff
changeset
|
3151 |
shows "f summable_on insert x A \<longleftrightarrow> f summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3152 |
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:
76999
diff
changeset
|
3153 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3154 |
lemma has_sum_insert: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3155 |
fixes f :: "'a \<Rightarrow> 'b :: topological_comm_monoid_add" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3156 |
assumes "x \<notin> A" and "(f has_sum S) A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3157 |
shows "(f has_sum (f x + S)) (insert x A)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3158 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3159 |
have "(f has_sum (f x + S)) ({x} \<union> A)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3160 |
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:
76999
diff
changeset
|
3161 |
thus ?thesis by simp |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3162 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3163 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3164 |
lemma infsum_insert: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3165 |
fixes f :: "_ \<Rightarrow> 'a :: {topological_comm_monoid_add, t2_space}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3166 |
assumes "f summable_on A" "a \<notin> A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3167 |
shows "infsum f (insert a A) = f a + infsum f A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3168 |
by (meson assms has_sum_insert infsumI summable_iff_has_sum_infsum) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3169 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3170 |
lemma has_sum_SigmaD: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3171 |
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:
76999
diff
changeset
|
3172 |
assumes sum1: "(f has_sum S) (Sigma A B)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3173 |
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:
76999
diff
changeset
|
3174 |
shows "(g has_sum S) A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3175 |
unfolding has_sum_def tendsto_def eventually_finite_subsets_at_top |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3176 |
proof (safe, goal_cases) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3177 |
case (1 X) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3178 |
with nhds_closed[of S X] obtain X' |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3179 |
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:
76999
diff
changeset
|
3180 |
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:
76999
diff
changeset
|
3181 |
by (auto simp: eventually_nhds) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3182 |
with sum1 obtain Y :: "('b \<times> 'c) set" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3183 |
where Y: "Y \<subseteq> Sigma A B" "finite Y" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3184 |
"\<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:
76999
diff
changeset
|
3185 |
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:
76999
diff
changeset
|
3186 |
define Y1 :: "'b set" where "Y1 = fst ` Y" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3187 |
from Y have Y1: "Y1 \<subseteq> A" by (auto simp: Y1_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3188 |
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:
76999
diff
changeset
|
3189 |
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:
76999
diff
changeset
|
3190 |
using that Y(1,2) unfolding Y2_def |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3191 |
by (force simp: image_iff intro: finite_subset[of _ "snd ` Y"])+ |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3192 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3193 |
show ?case |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3194 |
proof (rule exI[of _ Y1], safe, goal_cases) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3195 |
case (3 Z) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3196 |
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:
76999
diff
changeset
|
3197 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3198 |
have "sum g Z \<in> X'" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3199 |
proof (rule Lim_in_closed_set) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3200 |
show "closed X'" by fact |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3201 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3202 |
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:
76999
diff
changeset
|
3203 |
unfolding H_def |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3204 |
proof (intro tendsto_sum filterlim_INF') |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3205 |
fix x assume x: "x \<in> Z" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3206 |
with 3 have "x \<in> A" by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3207 |
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:
76999
diff
changeset
|
3208 |
by (simp add: has_sum_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3209 |
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:
76999
diff
changeset
|
3210 |
(filtercomap (\<lambda>p. p x) (finite_subsets_at_top (B x)))" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3211 |
by (rule filterlim_compose[OF _ filterlim_filtercomap]) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3212 |
qed auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3213 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3214 |
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:
76999
diff
changeset
|
3215 |
unfolding H_def |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3216 |
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:
76999
diff
changeset
|
3217 |
fix x assume x: "x \<in> Z" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3218 |
hence x': "x \<in> A" using 3 by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3219 |
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:
76999
diff
changeset
|
3220 |
(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:
76999
diff
changeset
|
3221 |
by (intro eventually_filtercomapI) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3222 |
(auto simp: eventually_finite_subsets_at_top intro: exI[of _ "Y2 x"]) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3223 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3224 |
fix h |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3225 |
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:
76999
diff
changeset
|
3226 |
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:
76999
diff
changeset
|
3227 |
using \<open>finite Z\<close> by (subst sum.Sigma) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3228 |
also have "\<dots> \<in> X''" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3229 |
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:
76999
diff
changeset
|
3230 |
also have "X'' \<subseteq> X'" by fact |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3231 |
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:
76999
diff
changeset
|
3232 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3233 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3234 |
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:
76999
diff
changeset
|
3235 |
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:
76999
diff
changeset
|
3236 |
unfolding H_def finite_subsets_at_top_def filtercomap_INF filtercomap_principal |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3237 |
by (simp add: INF_Sigma) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3238 |
also have "\<dots> \<noteq> bot" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3239 |
proof (rule INF_filter_not_bot, subst INF_principal_finite, goal_cases) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3240 |
case (2 X) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3241 |
define H' where |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3242 |
"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:
76999
diff
changeset
|
3243 |
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:
76999
diff
changeset
|
3244 |
by (force split: if_splits simp: H'_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3245 |
hence "H' \<noteq> {}" by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3246 |
thus "principal H' \<noteq> bot" by (simp add: principal_eq_bot_iff) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3247 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3248 |
finally show "H \<noteq> bot" . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3249 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3250 |
also have "X' \<subseteq> X" by fact |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3251 |
finally show "sum g Z \<in> X" . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3252 |
qed (insert Y(1,2), auto simp: Y1_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3253 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3254 |
|
82349
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3255 |
lemma has_sum_finite_iff: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3256 |
fixes S :: "'a :: {topological_comm_monoid_add,t2_space}" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3257 |
assumes "finite A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3258 |
shows "(f has_sum S) A \<longleftrightarrow> S = (\<Sum>x\<in>A. f x)" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3259 |
proof |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3260 |
assume "S = (\<Sum>x\<in>A. f x)" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3261 |
thus "(f has_sum S) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3262 |
by (intro has_sum_finiteI assms) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3263 |
next |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3264 |
assume "(f has_sum S) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3265 |
moreover have "(f has_sum (\<Sum>x\<in>A. f x)) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3266 |
by (intro has_sum_finiteI assms) auto |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3267 |
ultimately show "S = (\<Sum>x\<in>A. f x)" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3268 |
using has_sum_unique by blast |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3269 |
qed |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3270 |
|
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3271 |
lemma has_sum_finite_neutralI: |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3272 |
assumes "finite B" "B \<subseteq> A" "\<And>x. x \<in> A - B \<Longrightarrow> f x = 0" "c = (\<Sum>x\<in>B. f x)" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3273 |
shows "(f has_sum c) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3274 |
proof - |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3275 |
have "(f has_sum c) B" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3276 |
by (rule has_sum_finiteI) (use assms in auto) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3277 |
also have "?this \<longleftrightarrow> (f has_sum c) A" |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3278 |
by (intro has_sum_cong_neutral) (use assms in auto) |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3279 |
finally show ?thesis . |
a854ca7ca7d9
More migration from Theta_Functions_Library
paulson <lp15@cam.ac.uk>
parents:
81150
diff
changeset
|
3280 |
qed |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3281 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3282 |
lemma has_sum_SigmaI: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3283 |
fixes f :: "_ \<Rightarrow> 'a :: {topological_comm_monoid_add, t3_space}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3284 |
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:
76999
diff
changeset
|
3285 |
assumes g: "(g has_sum S) A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3286 |
assumes summable: "f summable_on Sigma A B" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3287 |
shows "(f has_sum S) (Sigma A B)" |
77388 | 3288 |
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:
76999
diff
changeset
|
3289 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3290 |
lemma summable_on_SigmaD1: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3291 |
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:
76999
diff
changeset
|
3292 |
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:
76999
diff
changeset
|
3293 |
assumes x: "x \<in> A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3294 |
shows "f x summable_on B x" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3295 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3296 |
have "(\<lambda>(x,y). f x y) summable_on Sigma {x} B" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3297 |
using f by (rule summable_on_subset) (use x in auto) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3298 |
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:
76999
diff
changeset
|
3299 |
by (intro summable_on_cong) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3300 |
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:
76999
diff
changeset
|
3301 |
by (intro summable_on_reindex [symmetric] inj_onI) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3302 |
also have "snd ` Sigma {x} B = B x" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3303 |
by (force simp: Sigma_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3304 |
finally show ?thesis . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3305 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3306 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3307 |
lemma has_sum_swap: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3308 |
"(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:
76999
diff
changeset
|
3309 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3310 |
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:
76999
diff
changeset
|
3311 |
by (rule bij_betwI[of _ _ _ "\<lambda>(x,y). (y,x)"]) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3312 |
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:
76999
diff
changeset
|
3313 |
by (simp add: case_prod_unfold) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3314 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3315 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3316 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3317 |
lemma summable_on_swap: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3318 |
"f summable_on (A \<times> B) \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) summable_on (B \<times> A)" |
77388 | 3319 |
by (metis has_sum_swap summable_on_def) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3320 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3321 |
lemma has_sum_cmult_right_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3322 |
fixes c :: "'a :: {topological_semigroup_mult, field}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3323 |
assumes "c \<noteq> 0" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3324 |
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:
76999
diff
changeset
|
3325 |
using has_sum_cmult_right[of f A "S/c" c] |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3326 |
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:
76999
diff
changeset
|
3327 |
by (auto simp: field_simps) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3328 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3329 |
lemma has_sum_cmult_left_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3330 |
fixes c :: "'a :: {topological_semigroup_mult, field}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3331 |
assumes "c \<noteq> 0" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3332 |
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:
76999
diff
changeset
|
3333 |
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:
76999
diff
changeset
|
3334 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3335 |
lemma finite_nonzero_values_imp_summable_on: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3336 |
assumes "finite {x\<in>X. f x \<noteq> 0}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3337 |
shows "f summable_on X" |
77388 | 3338 |
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:
76999
diff
changeset
|
3339 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3340 |
lemma summable_on_of_int_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3341 |
"(\<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:
76999
diff
changeset
|
3342 |
proof |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3343 |
assume "f summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3344 |
thus "(\<lambda>x. of_int (f x)) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3345 |
by (rule summable_on_homomorphism) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3346 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3347 |
assume "(\<lambda>x. of_int (f x) :: 'b) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3348 |
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:
76999
diff
changeset
|
3349 |
by (auto simp: summable_on_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3350 |
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:
76999
diff
changeset
|
3351 |
unfolding has_sum_def . |
77388 | 3352 |
moreover have "1/2 > (0 :: real)" |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3353 |
by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3354 |
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:
76999
diff
changeset
|
3355 |
(finite_subsets_at_top A)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3356 |
unfolding tendsto_iff by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3357 |
then obtain X where X: "finite X" "X \<subseteq> A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3358 |
"\<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:
76999
diff
changeset
|
3359 |
unfolding eventually_finite_subsets_at_top by metis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3360 |
|
77388 | 3361 |
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:
76999
diff
changeset
|
3362 |
proof - |
77388 | 3363 |
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:
76999
diff
changeset
|
3364 |
by (intro X) auto |
77388 | 3365 |
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:
76999
diff
changeset
|
3366 |
by (intro X that) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3367 |
ultimately have "dist (sum (\<lambda>x. of_int (f x)) X) (sum (\<lambda>x. of_int (f x) :: 'b) Y) < |
77388 | 3368 |
1/2 + 1/2" |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3369 |
using dist_triangle_less_add by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3370 |
thus ?thesis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3371 |
by (simp add: dist_norm flip: of_int_sum of_int_diff) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3372 |
qed |
77388 | 3373 |
then have "{x\<in>A. f x \<noteq> 0} \<subseteq> X" |
3374 |
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:
76999
diff
changeset
|
3375 |
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:
76999
diff
changeset
|
3376 |
using finite_subset by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3377 |
thus "f summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3378 |
by (rule finite_nonzero_values_imp_summable_on) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3379 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3380 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3381 |
lemma summable_on_of_nat_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3382 |
"(\<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:
76999
diff
changeset
|
3383 |
proof |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3384 |
assume "f summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3385 |
thus "(\<lambda>x. of_nat (f x) :: 'b) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3386 |
by (rule summable_on_homomorphism) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3387 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3388 |
assume "(\<lambda>x. of_nat (f x) :: 'b) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3389 |
hence "(\<lambda>x. of_int (int (f x)) :: 'b) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3390 |
by simp |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3391 |
also have "?this \<longleftrightarrow> (\<lambda>x. int (f x)) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3392 |
by (rule summable_on_of_int_iff) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3393 |
also have "\<dots> \<longleftrightarrow> f summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3394 |
by (simp add: summable_on_discrete_iff) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3395 |
finally show "f summable_on A" . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3396 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3397 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3398 |
lemma infsum_of_nat: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3399 |
"infsum (\<lambda>x::'a. of_nat (f x) :: 'b :: {real_normed_algebra_1}) A = of_nat (infsum f A)" |
77388 | 3400 |
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:
76999
diff
changeset
|
3401 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3402 |
lemma infsum_of_int: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3403 |
"infsum (\<lambda>x::'a. of_int (f x) :: 'b :: {real_normed_algebra_1}) A = of_int (infsum f A)" |
77388 | 3404 |
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:
76999
diff
changeset
|
3405 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3406 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3407 |
lemma summable_on_SigmaI: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3408 |
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:
76999
diff
changeset
|
3409 |
conditionally_complete_linorder}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3410 |
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:
76999
diff
changeset
|
3411 |
assumes g: "g summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3412 |
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:
76999
diff
changeset
|
3413 |
shows "f summable_on Sigma A B" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3414 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3415 |
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:
76999
diff
changeset
|
3416 |
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:
76999
diff
changeset
|
3417 |
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:
76999
diff
changeset
|
3418 |
using summable_on_imp_bounded_partial_sums[OF g] by blast |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3419 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3420 |
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:
76999
diff
changeset
|
3421 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3422 |
from C obtain X' where X': |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3423 |
"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:
76999
diff
changeset
|
3424 |
unfolding eventually_finite_subsets_at_top by metis |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3425 |
have "sum g X \<le> sum g (X \<union> X')" |
77388 | 3426 |
using X X' by (intro sum_mono2 g_nonneg) auto |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3427 |
also have "\<dots> \<le> C" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3428 |
using X X'(1,2) by (intro X'(3)) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3429 |
finally show ?thesis . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3430 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3431 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3432 |
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:
76999
diff
changeset
|
3433 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3434 |
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:
76999
diff
changeset
|
3435 |
have Y12: "Y = Sigma Y1 Y2" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3436 |
unfolding Y1_def Y2_def by force |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3437 |
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:
76999
diff
changeset
|
3438 |
using Y unfolding Y1_def Y2_def by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3439 |
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:
76999
diff
changeset
|
3440 |
using Y by (auto simp: Y1_def Y2_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3441 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3442 |
have "sum f Y = sum f (Sigma Y1 Y2)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3443 |
by (simp add: Y12) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3444 |
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:
76999
diff
changeset
|
3445 |
by (subst sum.Sigma) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3446 |
also have "\<dots> \<le> (\<Sum>x\<in>Y1. g x)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3447 |
proof (rule sum_mono) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3448 |
fix x assume x: "x \<in> Y1" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3449 |
show "(\<Sum>y\<in>Y2 x. f (x, y)) \<le> g x" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3450 |
proof (rule has_sum_mono') |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3451 |
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:
76999
diff
changeset
|
3452 |
using x by (intro has_sum_finite) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3453 |
show "((\<lambda>y. f (x, y)) has_sum g x) (B x)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3454 |
by (rule f) (use x Y12_subset in auto) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3455 |
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:
76999
diff
changeset
|
3456 |
using x that Y12_subset by (intro f_nonneg) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3457 |
qed (use Y12_subset in auto) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3458 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3459 |
also have "\<dots> \<le> C" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3460 |
using Y12_subset by (intro sum_g_le) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3461 |
finally show ?thesis . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3462 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3463 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3464 |
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:
76999
diff
changeset
|
3465 |
unfolding eventually_finite_subsets_at_top by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3466 |
thus ?thesis |
77388 | 3467 |
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:
76999
diff
changeset
|
3468 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3469 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3470 |
lemma summable_on_UnionI: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3471 |
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:
76999
diff
changeset
|
3472 |
conditionally_complete_linorder}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3473 |
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:
76999
diff
changeset
|
3474 |
assumes g: "g summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3475 |
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:
76999
diff
changeset
|
3476 |
assumes disj: "disjoint_family_on B A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3477 |
shows "f summable_on (\<Union>x\<in>A. B x)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3478 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3479 |
have "f \<circ> snd summable_on Sigma A B" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3480 |
using assms by (intro summable_on_SigmaI[where g = g]) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3481 |
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:
76999
diff
changeset
|
3482 |
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:
76999
diff
changeset
|
3483 |
also have "snd ` (Sigma A B) = (\<Union>x\<in>A. B x)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3484 |
by force |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3485 |
finally show ?thesis . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3486 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3487 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3488 |
lemma summable_on_SigmaD: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3489 |
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:
76999
diff
changeset
|
3490 |
assumes sum1: "f summable_on (Sigma A B)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3491 |
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:
76999
diff
changeset
|
3492 |
shows "(\<lambda>x. infsum (\<lambda>y. f (x, y)) (B x)) summable_on A" |
77388 | 3493 |
using assms unfolding summable_on_def |
3494 |
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:
76999
diff
changeset
|
3495 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3496 |
lemma summable_on_UnionD: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3497 |
fixes f :: "'a \<Rightarrow> 'c :: {topological_comm_monoid_add,t3_space}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3498 |
assumes sum1: "f summable_on (\<Union>x\<in>A. B x)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3499 |
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:
76999
diff
changeset
|
3500 |
assumes disj: "disjoint_family_on B A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3501 |
shows "(\<lambda>x. infsum f (B x)) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3502 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3503 |
have "(\<Union>x\<in>A. B x) = snd ` Sigma A B" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3504 |
by (force simp: Sigma_def) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3505 |
with sum1 have "f summable_on (snd ` Sigma A B)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3506 |
by simp |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3507 |
also have "?this \<longleftrightarrow> (f \<circ> snd) summable_on (Sigma A B)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3508 |
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:
76999
diff
changeset
|
3509 |
finally show "(\<lambda>x. infsum f (B x)) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3510 |
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:
76999
diff
changeset
|
3511 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3512 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3513 |
lemma summable_on_Union_iff: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3514 |
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:
76999
diff
changeset
|
3515 |
conditionally_complete_linorder, t3_space}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3516 |
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:
76999
diff
changeset
|
3517 |
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:
76999
diff
changeset
|
3518 |
assumes disj: "disjoint_family_on B A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3519 |
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:
76999
diff
changeset
|
3520 |
proof |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3521 |
assume "g summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3522 |
thus "f summable_on (\<Union>x\<in>A. B x)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3523 |
using summable_on_UnionI[of A f B g] assms by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3524 |
next |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3525 |
assume "f summable_on (\<Union>x\<in>A. B x)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3526 |
hence "(\<lambda>x. infsum f (B x)) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3527 |
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:
76999
diff
changeset
|
3528 |
also have "?this \<longleftrightarrow> g summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3529 |
using assms by (intro summable_on_cong) (auto simp: infsumI) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3530 |
finally show "g summable_on A" . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3531 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3532 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3533 |
lemma has_sum_Sigma': |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3534 |
fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3535 |
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:
76999
diff
changeset
|
3536 |
assumes summableAB: "(f has_sum a) (Sigma A B)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3537 |
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:
76999
diff
changeset
|
3538 |
shows "(b has_sum a) A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3539 |
by (intro has_sum_Sigma[OF _ assms] uniformly_continuous_add) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3540 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3541 |
lemma abs_summable_on_comparison_test': |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3542 |
assumes "g summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3543 |
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:
76999
diff
changeset
|
3544 |
shows "(\<lambda>x. norm (f x)) summable_on A" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3545 |
proof (rule Infinite_Sum.abs_summable_on_comparison_test) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3546 |
have "g summable_on A \<longleftrightarrow> (\<lambda>x. norm (g x)) summable_on A" |
77388 | 3547 |
by (metis summable_on_iff_abs_summable_on_real) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3548 |
with assms show "(\<lambda>x. norm (g x)) summable_on A" by blast |
77388 | 3549 |
qed (use assms in fastforce) |
77357
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3550 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3551 |
lemma has_sum_geometric_from_1: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3552 |
fixes z :: "'a :: {real_normed_field, banach}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3553 |
assumes "norm z < 1" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3554 |
shows "((\<lambda>n. z ^ n) has_sum (z / (1 - z))) {1..}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3555 |
proof - |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3556 |
have [simp]: "z \<noteq> 1" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3557 |
using assms by auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3558 |
have "(\<lambda>n. z ^ Suc n) sums (1 / (1 - z) - 1)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3559 |
using geometric_sums[of z] assms by (subst sums_Suc_iff) auto |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3560 |
also have "1 / (1 - z) - 1 = z / (1 - z)" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3561 |
by (auto simp: field_simps) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3562 |
finally have "(\<lambda>n. z ^ Suc n) sums (z / (1 - z))" . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3563 |
moreover have "summable (\<lambda>n. norm (z ^ Suc n))" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3564 |
using assms |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3565 |
by (subst summable_Suc_iff) (auto simp: norm_power intro!: summable_geometric) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3566 |
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:
76999
diff
changeset
|
3567 |
by (intro norm_summable_imp_has_sum) |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3568 |
also have "?this \<longleftrightarrow> ?thesis" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3569 |
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:
76999
diff
changeset
|
3570 |
finally show ?thesis . |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3571 |
qed |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3572 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3573 |
lemma has_sum_divide_const: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3574 |
fixes f :: "'a \<Rightarrow> 'b :: {topological_semigroup_mult, field, semiring_0}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3575 |
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:
76999
diff
changeset
|
3576 |
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:
76999
diff
changeset
|
3577 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3578 |
lemma has_sum_uminusI: |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3579 |
fixes f :: "'a \<Rightarrow> 'b :: {topological_semigroup_mult, ring_1}" |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3580 |
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:
76999
diff
changeset
|
3581 |
using has_sum_cmult_right[of f A S "-1"] by simp |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3582 |
|
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3583 |
end |
e65d8ee80811
New material contributed by Manuel
paulson <lp15@cam.ac.uk>
parents:
76999
diff
changeset
|
3584 |