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