src/HOL/Analysis/Infinite_Sum.thy
author paulson <lp15@cam.ac.uk>
Tue, 25 Jan 2022 14:13:33 +0000
changeset 75008 43b3d5318d72
parent 74791 227915e07891
child 75462 7448423e5dba
permissions -rw-r--r--
fixed dodgy intro! attributes
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.
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    20
See, e.g., Definition 4.11 in \cite{conway2013course}.
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
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    25
theory Infinite_Sum
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    26
  imports
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
    27
    Elementary_Topology
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    28
    "HOL-Library.Extended_Nonnegative_Real"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    29
    "HOL-Library.Complex_Order"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    30
begin
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    31
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    32
subsection \<open>Definition and syntax\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    33
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    34
definition has_sum :: \<open>('a \<Rightarrow> 'b :: {comm_monoid_add, topological_space}) \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool\<close> where
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    35
  \<open>has_sum f A x \<longleftrightarrow> (sum f \<longlongrightarrow> 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
    36
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    37
definition summable_on :: "('a \<Rightarrow> 'b::{comm_monoid_add, topological_space}) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "summable'_on" 46) where
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    38
  "f summable_on A \<longleftrightarrow> (\<exists>x. has_sum f A x)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    39
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    40
definition 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
    41
  "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
    42
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    43
abbreviation abs_summable_on :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "abs'_summable'_on" 46) where
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    44
  "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
    45
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    46
syntax (ASCII)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    47
  "_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add"  ("(3INFSUM (_/:_)./ _)" [0, 51, 10] 10)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    48
syntax
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    49
  "_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add"  ("(2\<Sum>\<^sub>\<infinity>(_/\<in>_)./ _)" [0, 51, 10] 10)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    50
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
    51
  "\<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
    52
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    53
syntax (ASCII)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    54
  "_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a"  ("(3INFSUM _./ _)" [0, 10] 10)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    55
syntax
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    56
  "_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a"  ("(2\<Sum>\<^sub>\<infinity>_./ _)" [0, 10] 10)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    57
translations
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    58
  "\<Sum>\<^sub>\<infinity>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
    59
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    60
syntax (ASCII)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    61
  "_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(3INFSUM _ |/ _./ _)" [0, 0, 10] 10)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    62
syntax
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    63
  "_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(2\<Sum>\<^sub>\<infinity>_ | (_)./ _)" [0, 0, 10] 10)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    64
translations
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    65
  "\<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
    66
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    67
print_translation \<open>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    68
let
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    69
  fun sum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    70
        if x <> y then raise Match
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    71
        else
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    72
          let
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    73
            val x' = Syntax_Trans.mark_bound_body (x, Tx);
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    74
            val t' = subst_bound (x', t);
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    75
            val P' = subst_bound (x', P);
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    76
          in
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    77
            Syntax.const @{syntax_const "_qinfsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    78
          end
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    79
    | sum_tr' _ = raise Match;
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    80
in [(@{const_syntax infsum}, K sum_tr')] end
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    81
\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    82
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    83
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
    84
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    85
lemma infsumI:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    86
  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
    87
  assumes \<open>has_sum 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
  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
    89
  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
    90
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    91
lemma infsum_eqI:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    92
  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
    93
  assumes \<open>x = y\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    94
  assumes \<open>has_sum 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
    95
  assumes \<open>has_sum g B y\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    96
  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
    97
  by (metis assms(1) assms(2) assms(3) 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
    98
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    99
lemma infsum_eqI':
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   100
  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
   101
  assumes \<open>\<And>x. has_sum f A x \<longleftrightarrow> has_sum g B x\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   102
  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
   103
  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
   104
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   105
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
   106
  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
   107
  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
   108
  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
   109
  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
   110
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   111
lemma summable_iff_has_sum_infsum: "f summable_on A \<longleftrightarrow> has_sum f A (infsum f A)"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   112
  using infsumI summable_on_def by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   113
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   114
lemma has_sum_infsum[simp]:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   115
  assumes \<open>f summable_on S\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   116
  shows \<open>has_sum f S (infsum f S)\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   117
  using assms by (auto simp: summable_on_def infsum_def has_sum_def tendsto_Lim)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   118
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   119
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
   120
  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
   121
  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
   122
  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
   123
  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
   124
  shows "has_sum f S x \<longleftrightarrow> has_sum g T x"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   125
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   126
  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
   127
      = 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
   128
  proof 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   129
    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
   130
    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
   131
      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
   132
    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
   133
    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
   134
      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
   135
    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
   136
    proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   137
      have \<open>P (sum f ((F\<inter>S) \<union> (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
   138
        apply (rule F0_P)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   139
        using \<open>F0 \<subseteq> S\<close>  \<open>finite F0\<close> that by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   140
      also have \<open>sum f ((F\<inter>S) \<union> (F0\<inter>S)) = 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
   141
        apply (rule sum.mono_neutral_cong)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   142
        using that \<open>finite F0\<close> F0'_def assms by auto
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   143
      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
   144
    qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   145
    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
   146
      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
   147
  next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   148
    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
   149
    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
   150
      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
   151
    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
   152
    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
   153
      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
   154
    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
   155
    proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   156
      have \<open>P (sum g ((F\<inter>T) \<union> (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
   157
        apply (rule F0_P)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   158
        using \<open>F0 \<subseteq> T\<close>  \<open>finite F0\<close> that by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   159
      also have \<open>sum g ((F\<inter>T) \<union> (F0\<inter>T)) = sum f F\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   160
        apply (rule sum.mono_neutral_cong)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   161
        using that \<open>finite F0\<close> F0'_def assms by auto
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   162
      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
   163
    qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   164
    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
   165
      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
   166
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   167
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   168
  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
   169
    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
   170
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   171
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   172
    by (simp add: has_sum_def)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   173
qed
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
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
   176
  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
   177
  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
   178
  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
   179
  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
   180
  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
   181
  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
   182
  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
   183
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   184
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
   185
  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
   186
  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
   187
  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
   188
  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
   189
  shows \<open>infsum f S = infsum g T\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   190
  apply (rule infsum_eqI')
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   191
  using assms by (rule 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
   192
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   193
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
   194
  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
   195
  shows "has_sum f A x \<longleftrightarrow> has_sum g A x"
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   196
  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
   197
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   198
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
   199
  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
   200
  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
   201
  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
   202
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   203
lemma infsum_cong:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   204
  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
   205
  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
   206
  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
   207
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   208
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
   209
  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
   210
  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
   211
  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
   212
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   213
  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
   214
    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
   215
  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
   216
  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
   217
    by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   218
  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
   219
      \<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
   220
  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
   221
    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
   222
    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
   223
      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
   224
      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
   225
    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
   226
    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
   227
      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
   228
    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
   229
      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
   230
    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
   231
      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
   232
      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
   233
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   234
  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
   235
    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
   236
  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
   237
    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
   238
    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
   239
  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
   240
    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
   241
    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
   242
        tendsto_compose_filtermap)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   243
  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
   244
    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
   245
  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
   246
    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
   247
    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
   248
    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
   249
  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
   250
    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
   251
      \<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
   252
  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
   253
    by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   254
  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
   255
    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
   256
  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
   257
    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
   258
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   259
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   260
lemma
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   261
  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
   262
  assumes \<open>has_sum f B b\<close> and \<open>has_sum f A a\<close> 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
   263
  shows has_sum_Diff: "has_sum f (B - A) (b - a)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   264
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   265
  have finite_subsets1:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   266
    "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
   267
  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
   268
    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
   269
    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
   270
      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
   271
      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
   272
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   273
    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
   274
      by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   275
    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
   276
      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
   277
      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
   278
    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
   279
      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
   280
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   281
  have finite_subsets2: 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   282
    "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
   283
    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
   284
      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
   285
      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
   286
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   287
  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
   288
    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
   289
  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
   290
    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
   291
  have "((\<lambda>F. sum 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
   292
  proof (subst asm_rl [of "(\<lambda>F. sum f (F\<inter>A)) = sum f o (\<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
   293
    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
   294
      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
   295
    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
   296
      unfolding o_def 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   297
      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
   298
        \<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
   299
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   300
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   301
  with limB have "((\<lambda>F. sum f F - sum f (F\<inter>A)) \<longlongrightarrow> b - a) (finite_subsets_at_top B)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   302
    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
   303
  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
   304
    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
   305
  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
   306
    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
   307
  hence "((\<lambda>F. sum f (F-A)) \<longlongrightarrow> b - a) (finite_subsets_at_top B)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   308
    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
   309
      \<open>((\<lambda>F. sum f F - sum f (F \<inter> A)) \<longlongrightarrow> b - a) (finite_subsets_at_top B)\<close> by fastforce
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   310
  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
   311
    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
   312
  hence limBA: "(sum f \<longlongrightarrow> b - a) (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
   313
    apply (rule tendsto_mono[rotated])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   314
    by (rule finite_subsets1)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   315
  thus ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   316
    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
   317
qed
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
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   320
lemma
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   321
  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
   322
  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
   323
  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
   324
  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
   325
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   326
lemma
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   327
  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
   328
  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
   329
  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
   330
  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
   331
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   332
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
   333
  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
   334
  (* Does this really require a linorder topology? (Instead of order topology.) *)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   335
  assumes \<open>has_sum f A a\<close> and "has_sum g B b"
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\<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
   337
  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
   338
  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
   339
  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
   340
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   341
  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
   342
  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
   343
    using assms(1,2) summable_on_def by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   344
  have \<open>has_sum f' (A\<union>B) a\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   345
    apply (subst has_sum_cong_neutral[where g=f and T=A])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   346
    by (auto simp: f'_def assms(1))
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   347
  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
   348
    by (meson has_sum_def)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   349
  have \<open>has_sum g' (A\<union>B) b\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   350
    apply (subst has_sum_cong_neutral[where g=g and T=B])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   351
    by (auto simp: g'_def assms(2))
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   352
  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
   353
    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
   354
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   355
  have *: \<open>\<forall>\<^sub>F x in finite_subsets_at_top (A \<union> B). sum f' x \<le> sum g' x\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   356
    apply (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
   357
    apply (rule sum_mono)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   358
    using assms by (auto simp: f'_def g'_def)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   359
  show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   360
    apply (rule tendsto_le)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   361
    using * g'_lim f'_lim by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   362
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   363
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   364
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
   365
  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
   366
  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
   367
  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
   368
  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
   369
  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
   370
  shows "infsum f A \<le> infsum g B"
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   371
  by (rule has_sum_mono_neutral[of f A _ g B _]) (use assms in \<open>auto intro: has_sum_infsum\<close>)
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   372
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   373
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
   374
  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
   375
  assumes "has_sum f A x" and "has_sum g A y"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   376
  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
   377
  shows "x \<le> y"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   378
  apply (rule 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
   379
  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
   380
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   381
lemma infsum_mono:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   382
  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
   383
  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
   384
  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
   385
  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
   386
  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
   387
  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
   388
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   389
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
   390
  assumes "finite F"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   391
  shows "has_sum f F (sum f F)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   392
  using assms
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   393
  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
   394
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   395
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
   396
  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
   397
  assumes "finite F"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   398
  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
   399
  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
   400
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   401
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
   402
  assumes "finite F"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   403
  shows "infsum f F = sum f F"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   404
  using assms by (auto intro: tendsto_Lim simp: finite_subsets_at_top_finite infsum_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
   405
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   406
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
   407
  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
   408
  assumes "has_sum f A x" and "\<epsilon> > 0"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   409
  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
   410
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   411
  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
   412
    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
   413
  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
   414
    using assms(2) by (rule tendstoD)
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   415
  thus ?thesis
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   416
    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
   417
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   418
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   419
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
   420
  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
   421
  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
   422
  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
   423
proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   424
  from assms have "has_sum f A (infsum f A)"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   425
    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
   426
  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
   427
    by (rule has_sum_finite_approximation)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   428
qed
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   429
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   430
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
   431
  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
   432
  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
   433
  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
   434
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   435
  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
   436
    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
   437
  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
   438
    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
   439
  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
   440
              (\<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
   441
  proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   442
    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
   443
    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
   444
      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
   445
    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
   446
      using lim
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   447
      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
   448
    
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   449
    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
   450
    proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   451
      from ev_P
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   452
      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
   453
        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
   454
      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
   455
      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
   456
        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
   457
      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
   458
        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
   459
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   460
      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
   461
      proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   462
        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
   463
          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
   464
        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
   465
          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
   466
        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
   467
          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
   468
        also have \<open>\<dots> < 2 * d\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   469
          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
   470
        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
   471
      qed
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   472
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   473
      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
   474
        by (rule dist_triangle3)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   475
      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
   476
        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
   477
      also have \<open>\<dots> \<le> e\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   478
        by (auto simp: d_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   479
      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
   480
    qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   481
    then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   482
      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
   483
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   484
  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
   485
    by (simp add: cauchy_filter_metric_filtermap)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   486
  then obtain L' where \<open>(sum f \<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
   487
    apply atomize_elim unfolding filterlim_def
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   488
    apply (rule complete_uniform[where S=UNIV, simplified, THEN iffD1, rule_format])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   489
      apply (auto simp add: filtermap_bot_iff)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   490
    by (meson Cauchy_convergent UNIV_I complete_def convergent_def)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   491
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   492
    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
   493
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   494
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   495
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
   496
  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
   497
  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
   498
  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
   499
  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
   500
  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
   501
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   502
  (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
   503
  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
   504
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   505
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
   506
  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
   507
    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
   508
  assumes "has_sum (\<lambda>x. norm (f x)) A n"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   509
  assumes "has_sum f A a"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   510
  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
   511
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   512
  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
   513
  proof-
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   514
    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
   515
      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
   516
      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
   517
    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
   518
      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
   519
      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
   520
    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
   521
      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
   522
    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
   523
      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
   524
    also have "\<dots> \<le> n + \<epsilon>"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   525
      apply (rule add_right_mono)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   526
      apply (rule has_sum_mono_neutral[where A=F and B=A and f=\<open>\<lambda>x. norm (f x)\<close> and g=\<open>\<lambda>x. norm (f x)\<close>])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   527
      using \<open>finite F\<close> \<open>F \<subseteq> A\<close> assms by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   528
    finally show ?thesis 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   529
      by assumption
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   530
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   531
  thus ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   532
    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
   533
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   534
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   535
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
   536
  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
   537
    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
   538
  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
   539
  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
   540
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
   541
  case True
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   542
  show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   543
    apply (rule norm_has_sum_bound[where A=A and f=f and a=\<open>infsum f A\<close> and n=\<open>infsum (\<lambda>x. norm (f x)) A\<close>])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   544
    using assms True
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   545
    by (metis 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
   546
next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   547
  case False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   548
  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
   549
    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
   550
  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
   551
    for X
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   552
    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
   553
  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
   554
  proof(rule ccontr)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   555
    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
   556
    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
   557
    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
   558
    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
   559
      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
   560
    moreover have "open S"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   561
    proof-
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   562
      have "closed {s::real. s \<ge> 0}"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   563
        using Elementary_Topology.closed_sequential_limits[where S = "{s::real. s \<ge> 0}"]
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   564
        by (metis Lim_bounded2 mem_Collect_eq)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   565
      moreover have "{s::real. s \<ge> 0} = UNIV - S"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   566
        unfolding S_def by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   567
      ultimately have "closed (UNIV - S)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   568
        by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   569
      thus ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   570
        by (simp add: Compl_eq_Diff_UNIV open_closed) 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   571
    qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   572
    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
   573
      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
   574
    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
   575
      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
   576
    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
   577
      by blast
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   578
    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
   579
      unfolding S_def by auto      
74642
8af77cb50c6d tuned proofs -- avoid z3, which is unavailable on arm64-linux;
wenzelm
parents: 74639
diff changeset
   580
    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
   581
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   582
  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
   583
    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
   584
  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
   585
    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
   586
    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
   587
  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
   588
    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
   589
  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
   590
    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
   591
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   592
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   593
lemma infsum_tendsto:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   594
  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
   595
  shows \<open>((\<lambda>F. sum f F) \<longlongrightarrow> infsum f S) (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
   596
  using assms by (simp flip: has_sum_def)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   597
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   598
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   599
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
   600
  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
   601
  shows \<open>has_sum 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
   602
  unfolding has_sum_def
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   603
  apply (subst tendsto_cong[where g=\<open>\<lambda>_. 0\<close>])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   604
   apply (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
   605
  using assms by (auto simp add: subset_iff)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   606
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   607
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
   608
  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
   609
  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
   610
  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
   611
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   612
lemma infsum_0:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   613
  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
   614
  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
   615
  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
   616
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   617
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
   618
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
   619
  by (simp_all add: infsum_0)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   620
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
   621
  by (simp_all add: 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
lemma has_sum_0_simp[simp]: \<open>has_sum (\<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
   623
  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
   624
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 has_sum_add:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   627
  fixes f g :: "'a \<Rightarrow> 'b::{topological_comm_monoid_add}"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   628
  assumes \<open>has_sum f A a\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   629
  assumes \<open>has_sum g A b\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   630
  shows \<open>has_sum (\<lambda>x. f x + g x) A (a + b)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   631
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   632
  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
   633
    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
   634
    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
   635
  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
   636
    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
   637
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   638
    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
   639
qed
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
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
   642
  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
   643
  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
   644
  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
   645
  shows \<open>(\<lambda>x. f x + g 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
   646
  by (metis (full_types) assms(1) assms(2) summable_on_def has_sum_add)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   647
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   648
lemma infsum_add:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   649
  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
   650
  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
   651
  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
   652
  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
   653
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   654
  have \<open>has_sum (\<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
   655
    by (simp add: assms(1) assms(2) has_sum_add)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   656
  then show ?thesis
74639
f831b6e589dc tuned proofs -- avoid z3, which is unavailable on arm64-linux;
wenzelm
parents: 74475
diff changeset
   657
    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
   658
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   659
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   660
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   661
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
   662
  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
   663
  assumes "has_sum f A a"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   664
  assumes "has_sum f B b"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   665
  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
   666
  shows \<open>has_sum f (A \<union> B) (a + b)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   667
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   668
  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
   669
    and \<open>fB x = (if x \<notin> A then f 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
   670
  have fA: \<open>has_sum fA (A \<union> B) a\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   671
    apply (subst has_sum_cong_neutral[where T=A and g=f])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   672
    using assms by (auto simp: fA_def)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   673
  have fB: \<open>has_sum fB (A \<union> B) b\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   674
    apply (subst has_sum_cong_neutral[where T=B and g=f])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   675
    using assms by (auto simp: fB_def)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   676
  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
   677
    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
   678
  show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   679
    unfolding fAB
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   680
    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
   681
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   682
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   683
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
   684
  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
   685
  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
   686
  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
   687
  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
   688
  shows \<open>f summable_on (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
   689
  by (meson assms(1) assms(2) disj summable_on_def 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
   690
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   691
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
   692
  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
   693
  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
   694
  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
   695
  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
   696
  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
   697
  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
   698
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   699
(* TODO move *)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   700
lemma (in uniform_space) cauchy_filter_complete_converges:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   701
  assumes "cauchy_filter F" "complete A" "F \<le> principal A" "F \<noteq> bot"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   702
  shows   "\<exists>c. F \<le> nhds c"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   703
  using assms unfolding complete_uniform by blast
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
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
   706
  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
   707
  \begin{itemize}
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   708
  \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
   709
      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
   710
      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
   711
      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
   712
      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
   713
  
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   714
  \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
   715
      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
   716
  \end{itemize}
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   717
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   718
  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
   719
  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
   720
  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
   721
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   722
lemma summable_on_subset:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   723
  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
   724
  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
   725
  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
   726
  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
   727
  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
   728
  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
   729
proof -
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   730
  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
   731
  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
   732
  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
   733
    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
   734
  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
   735
    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
   736
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   737
  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
   738
  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
   739
    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
   740
    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
   741
      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
   742
    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
   743
      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
   744
      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
   745
      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
   746
    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
   747
      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
   748
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   749
    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
   750
      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
   751
    then obtain P1 P2
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   752
      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
   753
        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
   754
        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
   755
      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
   756
      by auto
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   757
    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
   758
      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
   759
    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
   760
      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
   761
    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
   762
    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
   763
      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
   764
 
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   765
    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
   766
      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
   767
    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
   768
      have "P1 (sum f (F1' \<union> F0A))"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   769
        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
   770
      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
   771
        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
   772
    next
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   773
      have "P2 (sum f (F2' \<union> F0A))"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   774
        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
   775
      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
   776
        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
   777
    qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   778
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   779
    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
   780
      unfolding eventually_prod_filter
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   781
    proof (safe intro!: exI)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   782
      show "eventually (\<lambda>x. E' (x, sum f F0B)) (filtermap (sum f) (finite_subsets_at_top B))"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   783
       and "eventually (\<lambda>x. E' (sum f F0B, x)) (filtermap (sum f) (finite_subsets_at_top B))"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   784
        unfolding eventually_filtermap eventually_finite_subsets_at_top
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   785
        by (rule exI[of _ F0B]; use * in \<open>force simp: F0B_def\<close>)+
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   786
    next
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   787
      show "E (x, y)" if "E' (x, sum f F0B)" and "E' (sum f F0B, y)" for x y
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   788
        using E'E'E that by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   789
    qed
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   790
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   791
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   792
  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
   793
    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
   794
    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
   795
  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
   796
    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
   797
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   798
    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
   799
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   800
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   801
text \<open>A special case of @{thm [source] summable_on_subset} for Banach spaces with less premises.\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   802
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   803
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
   804
  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
   805
  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
   806
  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
   807
  shows \<open>f summable_on B\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   808
  by (rule summable_on_subset[OF _ _ assms])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   809
     (auto simp: complete_def convergent_def dest!: Cauchy_convergent)
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   810
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   811
lemma has_sum_empty[simp]: \<open>has_sum f {} 0\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   812
  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
   813
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   814
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
   815
  by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   816
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   817
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
   818
  by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   819
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   820
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
   821
  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
   822
  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
   823
  assumes conv: \<open>\<And>a. a \<in> A \<Longrightarrow> has_sum f (B a) (s a)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   824
  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
   825
  shows \<open>has_sum f (\<Union>a\<in>A. B a) (sum s A)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   826
  using assms
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   827
proof (insert finite conv disj, induction)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   828
  case empty
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   829
  then show ?case 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   830
    by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   831
next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   832
  case (insert x A)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   833
  have \<open>has_sum f (B x) (s x)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   834
    by (simp add: insert.prems)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   835
  moreover have IH: \<open>has_sum f (\<Union>a\<in>A. B a) (sum s A)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   836
    using insert by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   837
  ultimately have \<open>has_sum f (B x \<union> (\<Union>a\<in>A. B a)) (s x + sum s A)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   838
    apply (rule 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
   839
    using insert by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   840
  then show ?case
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   841
    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
   842
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   843
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   844
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   845
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
   846
  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
   847
  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
   848
  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
   849
  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
   850
  shows \<open>f summable_on (\<Union>a\<in>A. B a)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   851
  using finite conv disj apply induction by (auto intro!: 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
   852
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   853
lemma sum_infsum:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   854
  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
   855
  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
   856
  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
   857
  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
   858
  shows \<open>sum (\<lambda>a. infsum f (B a)) A = infsum f (\<Union>a\<in>A. B a)\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   859
  by (rule sym, rule infsumI)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   860
     (use sum_has_sum[of A f B \<open>\<lambda>a. infsum f (B 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
   861
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   862
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
   863
  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
   864
  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
   865
  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
   866
  (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
   867
  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
   868
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   869
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   870
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
   871
  fixes f :: \<open>'b :: {comm_monoid_add,topological_space} \<Rightarrow> 'c :: {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
   872
  assumes f_sum: \<open>\<And>F. finite F \<Longrightarrow> F \<subseteq> S \<Longrightarrow> sum (f o g) F = f (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
   873
      \<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
   874
  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
   875
    \<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
   876
  assumes infsum: \<open>has_sum g S x\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   877
  shows \<open>has_sum (f o g) S (f x)\<close> 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   878
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   879
  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
   880
    using infsum 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
   881
  then have \<open>((f o sum g) \<longlongrightarrow> f 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
   882
    apply (rule tendsto_compose_at)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   883
    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
   884
  then have \<open>(sum (f o g) \<longlongrightarrow> f x) (finite_subsets_at_top S)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   885
    apply (rule 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
   886
    using f_sum by fastforce
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   887
  then show \<open>has_sum (f o g) S (f x)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   888
    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
   889
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   890
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   891
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
   892
  fixes f :: \<open>'b :: {comm_monoid_add,topological_space} \<Rightarrow> 'c :: {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
   893
  assumes \<open>\<And>F. finite F \<Longrightarrow> F \<subseteq> S \<Longrightarrow> sum (f o g) F = f (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
   894
    \<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
   895
  assumes \<open>\<And>x. has_sum g S x \<Longrightarrow> 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
   896
    \<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
   897
  assumes \<open>g 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
   898
  shows \<open>(f o g) 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
   899
  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
   900
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   901
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
   902
  fixes f :: \<open>'b :: {comm_monoid_add,t2_space} \<Rightarrow> 'c :: {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
   903
  assumes f_sum: \<open>\<And>F. finite F \<Longrightarrow> F \<subseteq> S \<Longrightarrow> sum (f o g) F = f (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
   904
      \<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
   905
  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
   906
  assumes \<open>g 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
   907
  shows \<open>infsum (f o 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
   908
  using assms
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   909
  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
   910
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   911
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
   912
  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
   913
  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
   914
  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
   915
    \<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
   916
  assumes infsum: \<open>has_sum g S x\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   917
  shows \<open>has_sum (f o g) S (f x)\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   918
  using assms
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   919
  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
   920
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   921
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
   922
  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
   923
  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
   924
  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
   925
  assumes \<open>g 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
   926
  shows \<open>(f o g) 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
   927
  by (meson assms(1) assms(2) assms(3) summable_on_def has_sum_comm_additive has_sum_infsum isContD)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   928
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   929
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
   930
  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
   931
  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
   932
  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
   933
  assumes \<open>g 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
   934
  shows \<open>infsum (f o g) S = 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
   935
  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
   936
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   937
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
   938
  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
   939
  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
   940
  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
   941
  shows \<open>has_sum f 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
   942
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   943
  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
   944
  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
   945
    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
   946
    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
   947
      by (metis (mono_tags, lifting) Collect_cong Collect_empty_eq assms(2) empty_subsetI finite.emptyI less_cSUP_iff mem_Collect_eq)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   948
    show \<open>\<forall>\<^sub>F x in finite_subsets_at_top A. a < sum f x\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   949
      unfolding eventually_finite_subsets_at_top
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   950
    proof (rule exI[of _ F], safe)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   951
      fix Y assume Y: "finite Y" "F \<subseteq> Y" "Y \<subseteq> A"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   952
      have "a < sum f F"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   953
        by fact
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   954
      also have "\<dots> \<le> sum f Y"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   955
        using assms Y by (intro sum_mono2) auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   956
      finally show "a < sum f Y" .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   957
    qed (use \<open>finite F\<close> \<open>F \<subseteq> A\<close> in auto)
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   958
  next
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   959
    fix a assume *: \<open>(SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F) < a\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   960
    have \<open>sum f F < a\<close> if \<open>F\<subseteq>A\<close> and \<open>finite F\<close> for F
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   961
    proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   962
      have "sum f F \<le> (SUP F\<in>{F. finite F \<and> F\<subseteq>A}. sum f F)"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   963
        by (rule cSUP_upper) (use that assms(2) in \<open>auto simp: conj_commute\<close>)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   964
      also have "\<dots> < a"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   965
        by fact
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   966
      finally show ?thesis .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   967
    qed
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   968
    then show \<open>\<forall>\<^sub>F x in finite_subsets_at_top A. sum f x < a\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   969
      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
   970
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   971
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   972
    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
   973
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   974
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   975
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
   976
  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
   977
  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
   978
  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
   979
  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
   980
  using assms(1) assms(2) 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
   981
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   982
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
   983
  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
   984
  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
   985
  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
   986
  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
   987
  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
   988
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   989
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
   990
  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
   991
  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
   992
  shows \<open>has_sum 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
   993
  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
   994
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
   995
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
   996
  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
   997
  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
   998
  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
   999
  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
  1000
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1001
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
  1002
  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
  1003
  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
  1004
  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
  1005
  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
  1006
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1007
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
  1008
  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
  1009
  assumes "has_sum f M a"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1010
    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
  1011
  shows "a \<ge> 0"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1012
  by (metis (no_types, lifting) DiffD1 assms(1) assms(2) empty_iff has_sum_0 has_sum_mono_neutral order_refl)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1013
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1014
lemma infsum_nonneg:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1015
  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
  1016
  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
  1017
  shows "infsum f M \<ge> 0" (is "?lhs \<ge> _")
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1018
  apply (cases \<open>f summable_on M\<close>)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1019
   apply (metis assms infsum_0_simp summable_on_0_simp infsum_mono)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1020
  using assms by (auto simp add: infsum_not_exists)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1021
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1022
lemma has_sum_mono2:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1023
  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
  1024
  assumes "has_sum f A S" "has_sum f B S'" "A \<subseteq> B"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1025
  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
  1026
  shows   "S \<le> S'"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1027
proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1028
  have "has_sum f (B - A) (S' - S)"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1029
    by (rule has_sum_Diff) fact+
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1030
  hence "S' - S \<ge> 0"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1031
    by (rule has_sum_nonneg) (use assms(4) in auto)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1032
  thus ?thesis
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1033
    by (metis add_0 add_mono_thms_linordered_semiring(3) diff_add_cancel)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1034
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1035
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1036
lemma infsum_mono2:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1037
  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
  1038
  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
  1039
  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
  1040
  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
  1041
  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
  1042
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1043
lemma finite_sum_le_has_sum:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1044
  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
  1045
  assumes "has_sum f A S" "finite B" "B \<subseteq> A"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1046
  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
  1047
  shows   "sum f B \<le> S"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1048
proof (rule has_sum_mono2)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1049
  show "has_sum f A S"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1050
    by fact
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1051
  show "has_sum f B (sum f B)"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1052
    by (rule has_sum_finite) fact+
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1053
qed (use assms in auto)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1054
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1055
lemma finite_sum_le_infsum:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1056
  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
  1057
  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
  1058
  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
  1059
  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
  1060
  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
  1061
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1062
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
  1063
  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
  1064
  shows \<open>has_sum g (h ` A) x \<longleftrightarrow> has_sum (g \<circ> h) A x\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1065
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1066
  have \<open>has_sum g (h ` A) x \<longleftrightarrow> (sum g \<longlongrightarrow> x) (finite_subsets_at_top (h ` A))\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1067
    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
  1068
  also have \<open>\<dots> \<longleftrightarrow> ((\<lambda>F. sum g (h ` F)) \<longlongrightarrow> 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
  1069
    apply (subst filtermap_image_finite_subsets_at_top[symmetric])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1070
    using assms by (auto simp: filterlim_def filtermap_filtermap)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1071
  also have \<open>\<dots> \<longleftrightarrow> (sum (g \<circ> h) \<longlongrightarrow> 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
  1072
    apply (rule tendsto_cong)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1073
    apply (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
  1074
    apply (rule sum.reindex)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1075
    using assms subset_inj_on by blast
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1076
  also have \<open>\<dots> \<longleftrightarrow> has_sum (g \<circ> h) A x\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1077
    by (simp add: has_sum_def)
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1078
  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
  1079
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1080
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1081
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
  1082
  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
  1083
  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
  1084
  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
  1085
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1086
lemma infsum_reindex:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1087
  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
  1088
  shows \<open>infsum g (h ` A) = infsum (g \<circ> h) A\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1089
  by (metis (no_types, opaque_lifting) assms finite_subsets_at_top_neq_bot infsum_def 
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1090
        summable_on_reindex has_sum_def has_sum_infsum has_sum_reindex tendsto_Lim)
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1091
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1092
lemma summable_on_reindex_bij_betw:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1093
  assumes "bij_betw g A B"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1094
  shows   "(\<lambda>x. f (g x)) summable_on A \<longleftrightarrow> f summable_on B"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1095
proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1096
  thm summable_on_reindex
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1097
  have \<open>(\<lambda>x. f (g x)) summable_on A \<longleftrightarrow> f summable_on g ` A\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1098
    apply (rule summable_on_reindex[symmetric, unfolded o_def])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1099
    using assms bij_betw_imp_inj_on by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1100
  also have \<open>\<dots> \<longleftrightarrow> f summable_on B\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1101
    using assms bij_betw_imp_surj_on by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1102
  finally show ?thesis .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1103
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1104
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1105
lemma infsum_reindex_bij_betw:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1106
  assumes "bij_betw g A B"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1107
  shows   "infsum (\<lambda>x. f (g x)) A = infsum f B"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1108
proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1109
  have \<open>infsum (\<lambda>x. f (g x)) A = infsum f (g ` A)\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1110
    by (metis (mono_tags, lifting) assms bij_betw_imp_inj_on infsum_cong infsum_reindex o_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1111
  also have \<open>\<dots> = infsum f B\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1112
    using assms bij_betw_imp_surj_on by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1113
  finally show ?thesis .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1114
qed
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1115
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1116
lemma sum_uniformity:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1117
  assumes plus_cont: \<open>uniformly_continuous_on UNIV (\<lambda>(x::'b::{uniform_space,comm_monoid_add},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
  1118
  assumes \<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
  1119
  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
  1120
    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>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1121
proof (atomize_elim, insert \<open>eventually E uniformity\<close>, induction n arbitrary: E rule:nat_induct)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1122
  case 0
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1123
  then show ?case
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1124
    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
  1125
next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1126
  case (Suc n)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1127
  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
  1128
  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
  1129
    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
  1130
    apply atomize_elim
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1131
    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
  1132
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1133
  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
  1134
  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
  1135
    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
  1136
    by metis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1137
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1138
  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
  1139
  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
  1140
    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
  1141
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1142
  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
  1143
    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
  1144
    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
  1145
  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
  1146
    case True
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1147
    then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1148
      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
  1149
  next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1150
    case False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1151
    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
  1152
      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
  1153
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1154
    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
  1155
      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
  1156
    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
  1157
    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
  1158
      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
  1159
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1160
    from DM 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1161
    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
  1162
      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
  1163
    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
  1164
      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
  1165
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1166
    with D2_N
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1167
    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
  1168
      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
  1169
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1170
    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
  1171
      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
  1172
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1173
  with \<open>eventually D uniformity\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1174
  show ?case 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1175
    by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1176
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1177
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1178
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
  1179
  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
  1180
    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
  1181
  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
  1182
  assumes summableAB: "has_sum f (Sigma A B) a"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1183
  assumes summableB: \<open>\<And>x. x\<in>A \<Longrightarrow> has_sum (\<lambda>y. f (x, y)) (B 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
  1184
  shows "has_sum b A a"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1185
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1186
  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
  1187
    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
  1188
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1189
  from summableB
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1190
  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
  1191
    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
  1192
  from summableAB
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1193
  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
  1194
    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
  1195
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1196
  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
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1197
    apply (subst asm_rl[of \<open>{b| b. (a,b) \<in> H} = snd ` {ab. ab \<in> H \<and> fst ab = a}\<close>])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1198
    by (auto simp: image_iff that)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1199
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1200
  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
  1201
  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
  1202
    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
  1203
    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
  1204
    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
  1205
      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
  1206
    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
  1207
      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
  1208
      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
  1209
      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
  1210
    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
  1211
      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
  1212
    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
  1213
    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
  1214
    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
  1215
      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
  1216
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1217
    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
  1218
    proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1219
      define 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
  1220
      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
  1221
      proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1222
        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
  1223
          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>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1224
            for M' :: \<open>'a set\<close> and g g'
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1225
          apply (rule sum_uniformity[OF plus_cont \<open>eventually D uniformity\<close>, where n=\<open>card M\<close>])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1226
          by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1227
        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
  1228
          by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1229
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1230
        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
  1231
          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
  1232
        proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1233
          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
  1234
          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
  1235
            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
  1236
            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
  1237
          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
  1238
          ultimately show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1239
            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
  1240
            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
  1241
        qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1242
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1243
        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
  1244
        proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1245
          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
  1246
          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
  1247
            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
  1248
          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
  1249
            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
  1250
          then have *: \<open>(\<Sum>(a,b)\<in>H. f (a,b)) = (\<Sum>a\<in>M. \<Sum>b\<in>Ha' a. f (a,b))\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1251
            apply (subst sum.Sigma)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1252
            using \<open>finite M\<close> by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1253
          have \<open>D' (b a, sum (\<lambda>b. f (a,b)) (Ha' 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
  1254
            apply (rule D'_sum_Ha)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1255
            using that \<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
  1256
          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
  1257
            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
  1258
          with * show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1259
            by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1260
        qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1261
        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
  1262
          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
  1263
        ultimately show ?thesis
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1264
          unfolding FMB_def eventually_finite_subsets_at_top
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1265
          by (intro exI[of _ "Sigma M Ha"])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1266
             (use Ha_fin that(2,3) in \<open>fastforce intro!: finite_SigmaI\<close>)
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1267
      qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1268
      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
  1269
        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
  1270
      proof (rule exI[of _ G], safe)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1271
        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
  1272
        have "Y \<subseteq> Sigma A B"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1273
          using Y \<open>M \<subseteq> A\<close> by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1274
        thus "D (\<Sum>(a,b)\<in>Y. f (a, b), a)"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1275
          using G_sum[of Y] Y by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1276
      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
  1277
      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
  1278
        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
  1279
      then show \<open>E (sum b M, a)\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1280
        by (rule eventually_const[THEN iffD1, rotated]) (force simp: FMB_def)
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1281
    qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1282
    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
  1283
      using \<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
  1284
      by (auto intro!: exI[of _ \<open>fst ` G\<close>] simp add: FA_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
  1285
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1286
  then 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 (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
  1288
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1289
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1290
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
  1291
  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
  1292
    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
  1293
  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
  1294
  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
  1295
  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
  1296
  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
  1297
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1298
  from summableAB obtain a where a: \<open>has_sum (\<lambda>(x,y). f x y) (Sigma A B) a\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1299
    using has_sum_infsum by blast
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1300
  from summableB have b: \<open>\<And>x. x\<in>A \<Longrightarrow> has_sum (f x) (B x) (infsum (f 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
  1301
    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
  1302
  show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1303
    using plus_cont a b 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1304
    by (auto intro: has_sum_Sigma[where f=\<open>\<lambda>(x,y). f x y\<close>, simplified] 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
  1305
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1306
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1307
lemma infsum_Sigma:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1308
  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
  1309
    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
  1310
  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
  1311
  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
  1312
  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
  1313
  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
  1314
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1315
  from summableAB have a: \<open>has_sum f (Sigma A B) (infsum f (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
  1316
    using has_sum_infsum by blast
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1317
  from summableB have b: \<open>\<And>x. x\<in>A \<Longrightarrow> has_sum (\<lambda>y. f (x, y)) (B x) (infsum (\<lambda>y. f (x, y)) (B x))\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1318
    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
  1319
  show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1320
    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
  1321
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1322
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1323
lemma infsum_Sigma':
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1324
  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
  1325
    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
  1326
  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
  1327
  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
  1328
  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
  1329
  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
  1330
  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
  1331
  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
  1332
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1333
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
  1334
lemma
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1335
  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
  1336
    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
  1337
  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
  1338
  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
  1339
    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
  1340
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1341
  have [simp]: \<open>(f x) summable_on (B 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
  1342
  proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1343
    from assms
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1344
    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
  1345
      by (meson image_subset_iff summable_on_subset_banach mem_Sigma_iff that)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1346
    then have \<open>((\<lambda>(x,y). f x y) o Pair 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
  1347
      apply (rule_tac summable_on_reindex[THEN iffD1])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1348
      by (simp add: inj_on_def)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1349
    then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1350
      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
  1351
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1352
  show ?thesis1
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1353
    apply (rule infsum_Sigma')
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1354
    by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1355
  show ?thesis2
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1356
    apply (rule summable_on_Sigma)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1357
    by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1358
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1359
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1360
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
  1361
  fixes A :: "'a set" and B :: "'a \<Rightarrow> 'b set"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1362
    and f :: \<open>'a \<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
  1363
  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
  1364
  shows \<open>infsum (\<lambda>x. infsum (\<lambda>y. f (x,y)) (B x)) A = infsum f (Sigma A B)\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1365
  using assms
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1366
  by (subst infsum_Sigma'_banach) auto
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1367
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1368
lemma infsum_swap:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1369
  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
  1370
  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
  1371
  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
  1372
  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
  1373
  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
  1374
  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
  1375
  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
  1376
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1377
  have [simp]: \<open>(\<lambda>(x, y). f y x) summable_on (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
  1378
    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
  1379
    apply (subst summable_on_reindex)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1380
    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
  1381
  have \<open>infsum (\<lambda>x. infsum (\<lambda>y. f x y) B) A = infsum (\<lambda>(x,y). f x y) (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
  1382
    apply (subst infsum_Sigma)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1383
    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
  1384
  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
  1385
    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
  1386
    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
  1387
    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
  1388
  also have \<open>\<dots> = 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
  1389
    apply (subst infsum_Sigma)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1390
    using assms by auto
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1391
  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
  1392
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1393
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1394
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
  1395
  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
  1396
  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
  1397
  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
  1398
  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
  1399
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1400
  have [simp]: \<open>(\<lambda>(x, y). f y x) summable_on (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
  1401
    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
  1402
    apply (subst summable_on_reindex)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1403
    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
  1404
  have \<open>infsum (\<lambda>x. infsum (\<lambda>y. f x y) B) A = infsum (\<lambda>(x,y). f x y) (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
  1405
    apply (subst infsum_Sigma'_banach)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1406
    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
  1407
  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
  1408
    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
  1409
    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
  1410
    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
  1411
  also have \<open>\<dots> = 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
  1412
    apply (subst infsum_Sigma'_banach)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1413
    using assms by auto
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1414
  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
  1415
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1416
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1417
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
  1418
  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
  1419
  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
  1420
    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
  1421
    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
  1422
    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
  1423
  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
  1424
proof (rule ccontr)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1425
  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
  1426
  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
  1427
    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
  1428
  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
  1429
    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
  1430
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1431
  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
  1432
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1433
  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
  1434
    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
  1435
  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
  1436
    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
  1437
  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
  1438
    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
  1439
  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
  1440
    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
  1441
  finally show False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1442
    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
  1443
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1444
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1445
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
  1446
  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
  1447
  assumes "has_sum f A a" \<open>a \<le> 0\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1448
    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
  1449
    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
  1450
  shows "f x = 0"
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1451
  by (metis assms(1) assms(2) assms(4) infsumI nonneg_infsum_le_0D summable_on_def nneg)
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1452
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1453
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
  1454
  fixes f :: "'a \<Rightarrow> 'b :: {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
  1455
  assumes \<open>has_sum f A a\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1456
  shows "has_sum (\<lambda>x. f x * c) A (a * c)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1457
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1458
  from assms have \<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
  1459
    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
  1460
  then have \<open>((\<lambda>F. sum f F * c) \<longlongrightarrow> a * c) (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
  1461
    by (simp add: tendsto_mult_right)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1462
  then have \<open>(sum (\<lambda>x. f x * c) \<longlongrightarrow> a * c) (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
  1463
    apply (rule 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
  1464
    apply (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
  1465
    using sum_distrib_right by blast
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1466
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1467
    using infsumI 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
  1468
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1469
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1470
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
  1471
  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
  1472
  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
  1473
  shows "infsum (\<lambda>x. f x * c) A = infsum f A * c"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1474
proof (cases \<open>c=0\<close>)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1475
  case True
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1476
  then show ?thesis by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1477
next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1478
  case False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1479
  then have \<open>has_sum f 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
  1480
    by (simp add: assms)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1481
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1482
    by (auto intro!: infsumI 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
  1483
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1484
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1485
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
  1486
  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
  1487
  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
  1488
  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
  1489
  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
  1490
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1491
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
  1492
  fixes f :: "'a \<Rightarrow> 'b :: {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
  1493
  assumes \<open>has_sum f A a\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1494
  shows "has_sum (\<lambda>x. c * f x) A (c * a)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1495
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1496
  from assms have \<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
  1497
    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
  1498
  then have \<open>((\<lambda>F. c * sum f F) \<longlongrightarrow> c * 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
  1499
    by (simp add: tendsto_mult_left)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1500
  then have \<open>(sum (\<lambda>x. c * f x) \<longlongrightarrow> c * 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
  1501
    apply (rule 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
  1502
    apply (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
  1503
    using sum_distrib_left by blast
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1504
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1505
    using infsumI 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
  1506
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1507
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1508
lemma infsum_cmult_right:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1509
  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
  1510
  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
  1511
  shows \<open>infsum (\<lambda>x. c * f x) A = c * 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
  1512
proof (cases \<open>c=0\<close>)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1513
  case True
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1514
  then show ?thesis by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1515
next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1516
  case False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1517
  then have \<open>has_sum f 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
  1518
    by (simp add: assms)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1519
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1520
    by (auto intro!: infsumI 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
  1521
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1522
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1523
lemma summable_on_cmult_right:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1524
  fixes f :: "'a \<Rightarrow> 'b :: {t2_space, topological_semigroup_mult, semiring_0}"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1525
  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
  1526
  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
  1527
  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
  1528
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1529
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
  1530
  fixes f :: "'a \<Rightarrow> 'b :: {t2_space, topological_semigroup_mult, division_ring}"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1531
  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
  1532
  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
  1533
proof
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1534
  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
  1535
  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
  1536
    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
  1537
next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1538
  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
  1539
  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
  1540
    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
  1541
  then show \<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
  1542
    by (metis (no_types, lifting) assms summable_on_cong mult.assoc mult.right_neutral right_inverse)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1543
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1544
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1545
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
  1546
  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
  1547
  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
  1548
  shows "(\<lambda>x. c * f x) 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
  1549
proof
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1550
  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
  1551
  then show \<open>(\<lambda>x. c * f 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
  1552
    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
  1553
next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1554
  assume \<open>(\<lambda>x. c * f 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
  1555
  then have \<open>(\<lambda>x. inverse c * (c * f 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
  1556
    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
  1557
  then show \<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
  1558
    by (metis (no_types, lifting) assms summable_on_cong left_inverse mult.assoc mult.left_neutral)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1559
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1560
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1561
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
  1562
  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
  1563
  shows "infsum (\<lambda>x. f x * c) A = infsum f A * c"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1564
proof (cases \<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
  1565
  case True
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1566
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1567
    apply (rule_tac infsum_cmult_left) by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1568
next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1569
  case False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1570
  note asm = False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1571
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1572
  proof (cases \<open>c=0\<close>)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1573
    case True
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1574
    then show ?thesis by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1575
  next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1576
    case False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1577
    with asm have nex: \<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
  1578
      by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1579
    moreover have nex': \<open>\<not> (\<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
  1580
      using asm False apply (subst summable_on_cmult_left') by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1581
    ultimately show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1582
      unfolding infsum_def by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1583
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1584
qed
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 infsum_cmult_right':
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1587
  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
  1588
  shows "infsum (\<lambda>x. c * f x) A = c * infsum f A"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1589
proof (cases \<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
  1590
  case True
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1591
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1592
    apply (rule_tac infsum_cmult_right) by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1593
next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1594
  case False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1595
  note asm = False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1596
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1597
  proof (cases \<open>c=0\<close>)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1598
    case True
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1599
    then show ?thesis by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1600
  next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1601
    case False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1602
    with asm have nex: \<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
  1603
      by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1604
    moreover have nex': \<open>\<not> (\<lambda>x. c * f 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
  1605
      using asm False apply (subst summable_on_cmult_right') by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1606
    ultimately show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1607
      unfolding infsum_def by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1608
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1609
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1610
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1611
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1612
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
  1613
  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
  1614
  shows \<open>has_sum (\<lambda>_. c) F (of_nat (card F) * c)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1615
  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
  1616
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1617
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
  1618
  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
  1619
  shows \<open>infsum (\<lambda>_. c) F = of_nat (card F) * c\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1620
  apply (subst infsum_finite[OF assms]) by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1621
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1622
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
  1623
  \<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
  1624
       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
  1625
  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
  1626
  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
  1627
  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
  1628
proof (rule notI)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1629
  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
  1630
  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
  1631
    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
  1632
  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
  1633
    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
  1634
  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
  1635
  proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1636
    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
  1637
      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
  1638
    from assms
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1639
    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
  1640
      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
  1641
    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
  1642
    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
  1643
      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
  1644
    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
  1645
      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
  1646
      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
  1647
    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
  1648
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1649
  then show False
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1650
    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
  1651
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1652
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1653
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
  1654
  \<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
  1655
       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
  1656
  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
  1657
  shows \<open>infsum (\<lambda>_. c) A = of_nat (card A) * c\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1658
  apply (cases \<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
  1659
   apply simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1660
  apply (cases \<open>c = 0\<close>)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1661
   apply simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1662
  by (simp add: infsum_diverge_constant infsum_not_exists)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1663
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1664
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
  1665
  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
  1666
  shows \<open>has_sum (\<lambda>x. - f x) A a \<longleftrightarrow> has_sum f A (- a)\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1667
  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
  1668
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1669
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
  1670
  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
  1671
  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
  1672
  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
  1673
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1674
lemma infsum_uminus:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  1675
  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
  1676
  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
  1677
  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
  1678
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1679
lemma has_sum_le_finite_sums:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1680
  fixes a :: \<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
  1681
  assumes \<open>has_sum f A a\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1682
  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
  1683
  shows \<open>a \<le> b\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1684
proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1685
  from assms(1)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1686
  have 1: \<open>(sum f \<longlongrightarrow> a) (finite_subsets_at_top A)\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1687
    unfolding has_sum_def .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1688
  from assms(2)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1689
  have 2: \<open>\<forall>\<^sub>F F in finite_subsets_at_top A. sum f F \<le> b\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1690
    by (rule_tac eventually_finite_subsets_at_top_weakI)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1691
  show \<open>a \<le> b\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1692
    using _ _ 1 2
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1693
    apply (rule tendsto_le[where f=\<open>\<lambda>_. b\<close>])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1694
    by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1695
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1696
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1697
lemma infsum_le_finite_sums:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1698
  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
  1699
  assumes \<open>f summable_on A\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1700
  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
  1701
  shows \<open>infsum f A \<le> b\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1702
  by (meson assms(1) assms(2) has_sum_infsum has_sum_le_finite_sums)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1703
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1704
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1705
lemma summable_on_scaleR_left [intro]:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1706
  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
  1707
  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
  1708
  shows   "(\<lambda>x. f x *\<^sub>R c) summable_on A"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1709
  apply (cases \<open>c \<noteq> 0\<close>)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1710
   apply (subst asm_rl[of \<open>(\<lambda>x. f x *\<^sub>R c) = (\<lambda>y. y *\<^sub>R c) o f\<close>], simp add: o_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1711
   apply (rule summable_on_comm_additive)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1712
  using assms by (auto simp add: scaleR_left.additive_axioms)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1713
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1714
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1715
lemma summable_on_scaleR_right [intro]:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1716
  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
  1717
  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
  1718
  shows   "(\<lambda>x. c *\<^sub>R f x) summable_on A"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1719
  apply (cases \<open>c \<noteq> 0\<close>)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1720
   apply (subst asm_rl[of \<open>(\<lambda>x. c *\<^sub>R f x) = (\<lambda>y. c *\<^sub>R y) o f\<close>], simp add: o_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1721
   apply (rule summable_on_comm_additive)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1722
  using assms by (auto simp add: scaleR_right.additive_axioms)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1723
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1724
lemma infsum_scaleR_left:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1725
  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
  1726
  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
  1727
  shows   "infsum (\<lambda>x. f x *\<^sub>R c) A = infsum f A *\<^sub>R c"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1728
  apply (cases \<open>c \<noteq> 0\<close>)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1729
   apply (subst asm_rl[of \<open>(\<lambda>x. f x *\<^sub>R c) = (\<lambda>y. y *\<^sub>R c) o f\<close>], simp add: o_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1730
   apply (rule infsum_comm_additive)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1731
  using assms by (auto simp add: scaleR_left.additive_axioms)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1732
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1733
lemma infsum_scaleR_right:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1734
  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
  1735
  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
  1736
proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1737
  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
  1738
    by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1739
  then show ?thesis
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1740
  proof cases
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1741
    case summable
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1742
    then show ?thesis
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1743
      apply (subst asm_rl[of \<open>(\<lambda>x. c *\<^sub>R f x) = (\<lambda>y. c *\<^sub>R y) o f\<close>], simp add: o_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1744
      apply (rule infsum_comm_additive)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1745
      using summable by (auto simp add: scaleR_right.additive_axioms)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1746
  next
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1747
    case c0
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1748
    then show ?thesis by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1749
  next
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1750
    case not_summable
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1751
    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
  1752
    proof (rule notI)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1753
      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
  1754
      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
  1755
        using summable_on_scaleR_right by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1756
      then have \<open>f summable_on A\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1757
        using not_summable by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1758
      with not_summable show False
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1759
        by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1760
    qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1761
    then show ?thesis
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1762
      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
  1763
  qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1764
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1765
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1766
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1767
lemma infsum_Un_Int:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1768
  fixes f :: "'a \<Rightarrow> 'b::{topological_ab_group_add, t2_space}"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1769
  assumes [simp]: "f summable_on A - B" "f summable_on B - A" \<open>f summable_on A \<inter> B\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1770
  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
  1771
proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1772
  have [simp]: \<open>f summable_on A\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1773
    apply (subst asm_rl[of \<open>A = (A-B) \<union> (A\<inter>B)\<close>]) apply auto[1]
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1774
    apply (rule summable_on_Un_disjoint)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1775
    by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1776
  have \<open>infsum f (A \<union> B) = infsum f A + infsum f (B - A)\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1777
    apply (subst infsum_Un_disjoint[symmetric])
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
  moreover have \<open>infsum f (B - A \<union> A \<inter> B) = infsum f (B - A) + infsum f (A \<inter> B)\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1780
    by (rule infsum_Un_disjoint) auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1781
  moreover have "B - A \<union> A \<inter> B = B"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1782
    by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1783
  ultimately show ?thesis
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1784
    by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1785
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1786
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1787
lemma inj_combinator':
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1788
  assumes "x \<notin> F"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1789
  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
  1790
proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1791
  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
  1792
    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
  1793
  thus ?thesis
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1794
    by (simp add: o_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1795
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1796
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1797
lemma infsum_prod_PiE:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1798
  \<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
  1799
  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
  1800
  assumes finite: "finite A"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1801
  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
  1802
  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
  1803
  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
  1804
proof (use finite assms(2-) in induction)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1805
  case empty
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1806
  then show ?case 
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1807
    by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1808
next
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1809
  case (insert x F)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1810
  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
  1811
    unfolding PiE_insert_eq 
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1812
    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
  1813
  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
  1814
    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
  1815
  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
  1816
    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
  1817
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1818
  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
  1819
    using insert.prems(2) .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1820
  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
  1821
    by (simp only: pi)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1822
  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
  1823
               ((\<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
  1824
    using inj by (rule summable_on_reindex)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1825
  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
  1826
    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
  1827
  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
  1828
             (\<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
  1829
    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
  1830
  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
  1831
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1832
  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
  1833
    by (rule summable_on_Sigma_banach)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1834
  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
  1835
    apply (subst infsum_cmult_left[symmetric])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1836
    using insert.prems(1) by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1837
  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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1838
    apply (subst (asm) summable_on_cmult_right')
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1839
    using that by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1840
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1841
  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
  1842
     = (\<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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1843
    apply (subst pi)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1844
    apply (subst infsum_reindex)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1845
    using inj by (auto simp: o_def case_prod_unfold)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1846
  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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1847
    apply (subst prod.insert)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1848
    using insert by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1849
  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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1850
    apply (subst prod) by rule
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1851
  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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1852
    apply (subst infsum_Sigma_banach[symmetric])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1853
    using summable2 apply blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1854
    by fastforce
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1855
  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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1856
    apply (subst infsum_cmult_left')
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1857
    apply (subst infsum_cmult_right')
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1858
    by (rule refl)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1859
  also have \<open>\<dots> = (\<Prod>x\<in>insert x F. infsum (f x) (B x))\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1860
    apply (subst prod.insert)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1861
    using \<open>finite F\<close> \<open>x \<notin> F\<close> apply auto[2]
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1862
    apply (cases \<open>infsum (f x) (B x) = 0\<close>)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1863
     apply simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1864
    apply (subst insert.IH)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1865
      apply (simp add: insert.prems(1))
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1866
     apply (rule summable3)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1867
    by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1868
  finally show ?case
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1869
    by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1870
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1871
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1872
lemma infsum_prod_PiE_abs:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1873
  \<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
  1874
  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
  1875
  assumes finite: "finite A"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1876
  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
  1877
  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
  1878
proof (use finite assms(2) in induction)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1879
  case empty
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1880
  then show ?case 
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1881
    by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1882
next
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1883
  case (insert x F)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1884
  
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1885
  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
  1886
    unfolding PiE_insert_eq 
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1887
    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
  1888
  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
  1889
    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
  1890
  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
  1891
    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
  1892
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1893
  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
  1894
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1895
  have *: \<open>(\<Sum>p\<in>P. norm (\<Prod>x\<in>F. f x (p x))) \<le> prod s F\<close> 
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1896
    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
  1897
      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
  1898
  proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1899
    define B' where \<open>B' x = {p x| p. p\<in>P}\<close> for x
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1900
    have [simp]: \<open>finite (B' x)\<close> for x
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1901
      using that by (auto simp: B'_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1902
    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
  1903
      by (simp add: finite_PiE)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1904
    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
  1905
      using that by (auto simp: B'_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1906
    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
  1907
      unfolding B'_def using P that 
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1908
      by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1909
    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
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1910
      apply (simp_all add: s_def flip: infsum_finite)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1911
      apply (rule infsum_mono_neutral)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1912
      using that sum B'B by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1913
    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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1914
      apply (rule sum_mono2)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1915
      by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1916
    also have \<open>\<dots> = (\<Sum>p\<in>Pi\<^sub>E F B'. \<Prod>x\<in>F. norm (f x (p x)))\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1917
      apply (subst prod_norm[symmetric])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1918
      by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1919
    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
  1920
    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
  1921
      case empty
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1922
      then show ?case by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1923
    next
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1924
      case (insert x F)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1925
      have aux: \<open>a = b \<Longrightarrow> c * a = c * b\<close> for a b c :: real
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1926
        by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1927
      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
  1928
        by (rule inj_combinator') (use insert.hyps in auto)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1929
      have \<open>(\<Sum>p\<in>Pi\<^sub>E (insert x F) B'. \<Prod>x\<in>insert x F. norm (f x (p x)))
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1930
         =  (\<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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1931
        apply (subst pi)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1932
        apply (subst sum.reindex)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1933
        using inj by (auto simp: case_prod_unfold)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1934
      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 := y)) x'))))\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1935
        apply (subst prod.insert)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1936
        using insert.hyps by (auto simp: case_prod_unfold)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1937
      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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1938
        apply (rule sum.cong)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1939
         apply blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1940
        unfolding case_prod_unfold
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1941
        apply (rule aux)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1942
        apply (rule prod.cong)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1943
        using insert.hyps(2) by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1944
      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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1945
        apply (subst sum_product)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1946
        apply (subst sum.swap)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1947
        apply (subst sum.cartesian_product)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1948
        by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1949
      also have \<open>\<dots> = (\<Sum>y\<in>B' x. norm (f x y)) * (\<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
  1950
        by (simp add: insert.IH)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1951
      also have \<open>\<dots> = (\<Prod>x\<in>insert x 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
  1952
        using insert.hyps(1) insert.hyps(2) by force
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1953
      finally show ?case .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1954
    qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1955
    also have \<open>\<dots> = (\<Prod>x\<in>F. \<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
  1956
      by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1957
    also have \<open>\<dots> \<le> (\<Prod>x\<in>F. s x)\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1958
      apply (rule prod_mono)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1959
      apply auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1960
      apply (simp add: sum_nonneg)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1961
      using s_bound by presburger
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1962
    finally show ?thesis .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1963
  qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1964
  have \<open>(\<lambda>g. \<Prod>x\<in>insert x F. f x (g x)) abs_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
  1965
    apply (rule nonneg_bdd_above_summable_on)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1966
     apply (simp; fail)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1967
    apply (rule bdd_aboveI[where M=\<open>\<Prod>x'\<in>insert x F. s x'\<close>])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1968
    using * insert.hyps insert.prems by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1969
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1970
  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
  1971
    by (simp only: pi)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1972
  also have "(\<lambda>g. \<Prod>x\<in>insert x F. f x (g x)) abs_summable_on \<dots> \<longleftrightarrow>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1973
               ((\<lambda>g. \<Prod>x\<in>insert x F. f x (g x)) \<circ> (\<lambda>(g,y). g(x:=y))) abs_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
  1974
    using inj by (subst summable_on_reindex) (auto simp: o_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1975
  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
  1976
    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
  1977
  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
  1978
             (\<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
  1979
    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
  1980
  finally have summable2: \<open>(\<lambda>(p, y). f x y * (\<Prod>x'\<in>F. f x' (p x'))) abs_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
  1981
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1982
  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
  1983
     = (\<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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1984
    apply (subst pi)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1985
    apply (subst infsum_reindex)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1986
    using inj by (auto simp: o_def case_prod_unfold)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1987
  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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1988
    apply (subst prod.insert)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1989
    using insert by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1990
  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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1991
    apply (subst prod) by rule
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1992
  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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1993
    apply (subst infsum_Sigma_banach[symmetric])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1994
    using summable2 abs_summable_summable apply blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1995
    by fastforce
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1996
  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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1997
    apply (subst infsum_cmult_left')
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1998
    apply (subst infsum_cmult_right')
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  1999
    by (rule refl)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2000
  also have \<open>\<dots> = (\<Prod>x\<in>insert x F. infsum (f x) (B x))\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2001
    apply (subst prod.insert)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2002
    using \<open>finite F\<close> \<open>x \<notin> F\<close> apply auto[2]
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2003
    apply (cases \<open>infsum (f x) (B x) = 0\<close>)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2004
     apply (simp; fail)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2005
    apply (subst insert.IH)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2006
      apply (auto simp add: insert.prems(1))
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2007
    done
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2008
  finally show ?case
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2009
    by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2010
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2011
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2012
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2013
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2014
subsection \<open>Absolute convergence\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2015
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2016
lemma abs_summable_countable:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2017
  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
  2018
  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
  2019
proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2020
  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
  2021
  proof (rule ccontr)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2022
    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
  2023
    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
  2024
    proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2025
      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
  2026
        by (meson real_arch_simple)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2027
      from *
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2028
      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
  2029
        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
  2030
      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
  2031
        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
  2032
      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
  2033
        by (simp add: cardF that)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2034
      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
  2035
        by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2036
      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
  2037
        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
  2038
      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
  2039
        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
  2040
      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
  2041
        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
  2042
      finally show ?thesis .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2043
    qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2044
    then show False
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2045
      by (meson gt_ex linorder_not_less)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2046
  qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2047
  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
  2048
    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
  2049
  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
  2050
  proof safe
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2051
    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
  2052
    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
  2053
    have "i \<ge> 1"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2054
      by (simp add: i_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2055
    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
  2056
      unfolding i_def by linarith
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2057
    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
  2058
      by (auto simp: divide_simps mult_ac)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2059
    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
  2060
      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
  2061
  qed auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2062
  finally show ?thesis .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2063
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2064
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2065
(* 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
  2066
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
  2067
  fixes f :: \<open>'a \<Rightarrow> real\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2068
  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
  2069
proof (rule iffI)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2070
  assume \<open>f summable_on A\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2071
  define n A\<^sub>p A\<^sub>n
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2072
    where \<open>n 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
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2073
  have [simp]: \<open>A\<^sub>p \<union> A\<^sub>n = A\<close> \<open>A\<^sub>p \<inter> A\<^sub>n = {}\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2074
    by (auto simp: A\<^sub>p_def A\<^sub>n_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2075
  from \<open>f summable_on A\<close> have [simp]: \<open>f summable_on A\<^sub>p\<close> \<open>f summable_on A\<^sub>n\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2076
    using A\<^sub>p_def A\<^sub>n_def summable_on_subset_banach by fastforce+
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2077
  then have [simp]: \<open>n summable_on A\<^sub>p\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2078
    apply (subst summable_on_cong[where g=f])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2079
    by (simp_all add: A\<^sub>p_def n_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2080
  moreover have [simp]: \<open>n summable_on A\<^sub>n\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2081
    apply (subst summable_on_cong[where g=\<open>\<lambda>x. - f x\<close>])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2082
     apply (simp add: A\<^sub>n_def n_def[abs_def])
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2083
    by (simp add: summable_on_uminus)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2084
  ultimately have [simp]: \<open>n summable_on (A\<^sub>p \<union> A\<^sub>n)\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2085
    apply (rule summable_on_Un_disjoint) by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2086
  then show \<open>n summable_on A\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2087
    by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2088
next
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2089
  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
  2090
    using abs_summable_summable by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2091
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2092
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2093
lemma abs_summable_on_Sigma_iff:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2094
  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
  2095
             (\<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
  2096
             ((\<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
  2097
proof (intro iffI conjI ballI)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2098
  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
  2099
  then have \<open>(\<lambda>x. infsum (\<lambda>y. norm (f (x,y))) (B x)) summable_on A\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2100
    apply (rule_tac summable_on_Sigma_banach)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2101
    by (auto simp: case_prod_unfold)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2102
  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
  2103
    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
  2104
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2105
  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
  2106
  proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2107
    from asm have \<open>f abs_summable_on Pair x ` B x\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2108
      apply (rule summable_on_subset_banach)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2109
      using that by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2110
    then show ?thesis
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2111
      apply (subst (asm) summable_on_reindex)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2112
      by (auto simp: o_def inj_on_def)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2113
  qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2114
next
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2115
  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
  2116
    (\<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
  2117
  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
  2118
    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
  2119
  proof -
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2120
    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
  2121
      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
  2122
    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
  2123
      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
  2124
    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>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2125
      apply (subst sum.Sigma)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2126
      by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2127
    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
  2128
      apply (subst infsum_finite)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2129
      by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2130
    also have \<open>\<dots> \<le> (\<Sum>\<^sub>\<infinity>x\<in>fst ` F. \<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
  2131
      apply (rule infsum_mono)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2132
        apply (simp; fail)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2133
       apply (simp; fail)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2134
      apply (rule infsum_mono_neutral)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2135
      using asm that(1) by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2136
    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
  2137
      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
  2138
    finally show ?thesis .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2139
  qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2140
  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
  2141
    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
  2142
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2143
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2144
lemma abs_summable_on_comparison_test:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2145
  assumes "g abs_summable_on A"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2146
  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
  2147
  shows   "f abs_summable_on A"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2148
proof (rule nonneg_bdd_above_summable_on)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2149
  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
  2150
  proof (rule bdd_aboveI2)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2151
    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
  2152
    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
  2153
      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
  2154
    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
  2155
      using F by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2156
    also have \<open>\<dots> \<le> infsum (\<lambda>x. norm (g x)) A\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2157
    proof (rule infsum_mono_neutral)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2158
      show "g abs_summable_on F"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2159
        by (rule summable_on_subset_banach[OF assms(1)]) (use F in auto)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2160
    qed (use F assms in auto)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2161
    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
  2162
  qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2163
qed auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2164
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2165
lemma abs_summable_iff_bdd_above:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2166
  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
  2167
  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
  2168
proof (rule iffI)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2169
  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
  2170
  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
  2171
  proof (rule bdd_aboveI2)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2172
    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
  2173
    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
  2174
      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
  2175
  qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2176
next
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2177
  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
  2178
  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
  2179
    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
  2180
qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2181
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2182
lemma abs_summable_product:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2183
  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
  2184
  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
  2185
    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
  2186
  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
  2187
proof (rule nonneg_bdd_above_summable_on)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2188
  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
  2189
  proof (rule bdd_aboveI2)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2190
    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
  2191
    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
  2192
      by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2193
  
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2194
    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))"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2195
      apply (rule infsum_mono_neutral)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2196
      using b4 r1 x2_sum by auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2197
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2198
    have "norm (x i * y i) \<le> norm (x i * x i) + norm (y i * y i)" for i
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2199
      unfolding norm_mult by (smt mult_left_mono mult_nonneg_nonneg mult_right_mono norm_ge_zero)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2200
    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
  2201
      by (simp add: sum_mono)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2202
    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
  2203
      by (simp add: sum.distrib)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2204
    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
  2205
      by (simp add: \<open>finite F\<close>)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2206
    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
  2207
      using F assms
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2208
      by (intro add_mono infsum_mono2) auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2209
    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
  2210
      by simp
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2211
  qed
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2212
qed auto
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2213
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2214
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
  2215
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2216
lemma summable_on_ennreal[simp]: \<open>(f::_ \<Rightarrow> ennreal) summable_on S\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2217
  by (rule nonneg_summable_on_complete) simp
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2218
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2219
lemma summable_on_enat[simp]: \<open>(f::_ \<Rightarrow> enat) summable_on S\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2220
  by (rule nonneg_summable_on_complete) simp
74475
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_superconst_infinite_ennreal:
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> ennreal\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2224
  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
  2225
  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
  2226
  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
  2227
  shows "has_sum f S \<infinity>"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2228
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2229
  have \<open>(sum f \<longlongrightarrow> \<infinity>) (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
  2230
  proof (rule order_tendstoI[rotated], simp)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2231
    fix y :: ennreal assume \<open>y < \<infinity>\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2232
    then have \<open>y / b < \<infinity>\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2233
      by (metis b ennreal_divide_eq_top_iff gr_implies_not_zero infinity_ennreal_def top.not_eq_extremum)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2234
    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
  2235
      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
  2236
      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
  2237
    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
  2238
    proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2239
      have \<open>y < b * card F\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2240
        by (metis \<open>y < \<infinity>\<close> b cardF divide_less_ennreal ennreal_mult_eq_top_iff gr_implies_not_zero infinity_ennreal_def mult.commute top.not_eq_extremum)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2241
      also have \<open>\<dots> \<le> b * card Y\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2242
        by (meson b card_mono less_imp_le mult_left_mono of_nat_le_iff that(1) that(2))
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2243
      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
  2244
        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
  2245
      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
  2246
        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
  2247
      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
  2248
    qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2249
    ultimately show \<open>\<forall>\<^sub>F x in finite_subsets_at_top S. y < sum f x\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2250
      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
  2251
      by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2252
  qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2253
  then show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2254
    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
  2255
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2256
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2257
lemma infsum_superconst_infinite_ennreal:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2258
  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
  2259
  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
  2260
  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
  2261
  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
  2262
  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
  2263
  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
  2264
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2265
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
  2266
  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
  2267
  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
  2268
  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
  2269
  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
  2270
  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
  2271
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2272
  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
  2273
    using b by blast
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2274
  have "0 < e2ennreal b"
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2275
    using b' b
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2276
    by (metis dual_order.refl enn2ereal_e2ennreal gr_zeroI order_less_le zero_ennreal.abs_eq)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2277
  hence *: \<open>infsum (e2ennreal o f) S = \<infinity>\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2278
    using assms b'
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2279
    by (intro infsum_superconst_infinite_ennreal[where b=b']) (auto intro!: e2ennreal_mono)
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2280
  have \<open>infsum f S = infsum (enn2ereal o (e2ennreal o f)) S\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2281
    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
  2282
  also have \<open>\<dots> = enn2ereal \<infinity>\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2283
    apply (subst 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
  2284
    using * by (auto simp: continuous_at_enn2ereal)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2285
  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
  2286
    by simp
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2287
  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
  2288
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2289
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2290
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
  2291
  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
  2292
  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
  2293
  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
  2294
  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
  2295
  shows "has_sum f S \<infinity>"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2296
  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
  2297
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2298
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
  2299
  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
  2300
  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
  2301
  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
  2302
  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
  2303
  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
  2304
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2305
  have \<open>ennreal_of_enat (infsum f S) = infsum (ennreal_of_enat o f) S\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2306
    apply (rule infsum_comm_additive_general[symmetric])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2307
    by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2308
  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
  2309
    by (metis assms(3) b comp_apply ennreal_of_enat_0 ennreal_of_enat_inj ennreal_of_enat_le_iff geqb infsum_superconst_infinite_ennreal not_gr_zero)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2310
  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
  2311
    by simp
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2312
  finally show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2313
    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
  2314
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2315
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2316
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
  2317
  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
  2318
  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
  2319
  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
  2320
  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
  2321
  shows "has_sum f S \<infinity>"
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2322
  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
  2323
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2324
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
  2325
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2326
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
  2327
  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
  2328
  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
  2329
    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
  2330
  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
  2331
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2332
  have \<open>ennreal (infsum f A) = infsum (ennreal o f) A\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2333
    apply (rule infsum_comm_additive_general[symmetric])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2334
    apply (subst sum_ennreal[symmetric])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2335
    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
  2336
  also have \<open>\<dots> = (SUP F\<in>{F. finite F \<and> F \<subseteq> A}. (ennreal (sum f F)))\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2337
    apply (subst nonneg_infsum_complete, simp)
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2338
    apply (rule SUP_cong, blast)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2339
    apply (subst sum_ennreal[symmetric])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2340
    using fnn by auto
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2341
  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
  2342
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2343
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2344
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
  2345
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2346
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
  2347
  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
  2348
  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
  2349
    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
  2350
  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
  2351
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2352
  have \<open>ereal (infsum f A) = infsum (ereal o f) A\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2353
    apply (rule infsum_comm_additive_general[symmetric])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2354
    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
  2355
  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
  2356
    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
  2357
  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
  2358
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2359
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2360
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2361
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
  2362
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2363
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
  2364
      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
  2365
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2366
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
  2367
  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
  2368
  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
  2369
    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
  2370
  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
  2371
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2372
  have "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
  2373
    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
  2374
  also have "\<dots> = ereal (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
  2375
  proof (subst ereal_SUP)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2376
    show "\<bar>SUP a\<in>{F. finite F \<and> F \<subseteq> A}. ereal (sum f a)\<bar> \<noteq> \<infinity>"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2377
      using calculation by fastforce      
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2378
    show "(SUP F\<in>{F. finite F \<and> F \<subseteq> A}. ereal (sum f F)) = (SUP a\<in>{F. finite F \<and> F \<subseteq> A}. ereal (sum f a))"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2379
      by simp      
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
  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
  2382
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2383
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2384
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2385
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
  2386
  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
  2387
  assumes "f summable_on A" and "\<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
  2388
  shows "has_sum 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
  2389
  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
  2390
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2391
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
  2392
  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
  2393
  assumes \<open>f summable_on A\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2394
  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
  2395
  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
  2396
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2397
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
  2398
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2399
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
  2400
  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
  2401
  shows \<open>has_sum (\<lambda>x. cnj (f x)) M (cnj a) \<longleftrightarrow> has_sum f M a\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2402
  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
  2403
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2404
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
  2405
  "(\<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
  2406
  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
  2407
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2408
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
  2409
  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
  2410
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2411
lemma infsum_Re:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2412
  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
  2413
  shows "infsum (\<lambda>x. Re (f x)) M = Re (infsum f M)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2414
  apply (rule infsum_comm_additive[where f=Re, unfolded o_def])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2415
  using assms by (auto intro!: additive.intro)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2416
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2417
lemma has_sum_Re:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2418
  assumes "has_sum f M a"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2419
  shows "has_sum (\<lambda>x. Re (f x)) M (Re a)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2420
  apply (rule has_sum_comm_additive[where f=Re, unfolded o_def])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2421
  using assms by (auto intro!: additive.intro tendsto_Re)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2422
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2423
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
  2424
  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
  2425
  shows "(\<lambda>x. Re (f x)) summable_on M"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2426
  apply (rule summable_on_comm_additive[where f=Re, unfolded o_def])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2427
  using assms by (auto intro!: additive.intro)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2428
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2429
lemma infsum_Im: 
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2430
  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
  2431
  shows "infsum (\<lambda>x. Im (f x)) M = Im (infsum f M)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2432
  apply (rule infsum_comm_additive[where f=Im, unfolded o_def])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2433
  using assms by (auto intro!: additive.intro)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2434
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2435
lemma has_sum_Im:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2436
  assumes "has_sum f M a"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2437
  shows "has_sum (\<lambda>x. Im (f x)) M (Im a)"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2438
  apply (rule has_sum_comm_additive[where f=Im, unfolded o_def])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2439
  using assms by (auto intro!: additive.intro tendsto_Im)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2440
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2441
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
  2442
  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
  2443
  shows "(\<lambda>x. Im (f x)) summable_on M"
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2444
  apply (rule summable_on_comm_additive[where f=Im, unfolded o_def])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2445
  using assms by (auto intro!: additive.intro)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2446
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2447
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
  2448
  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
  2449
  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
  2450
    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
  2451
    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
  2452
    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
  2453
  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
  2454
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2455
  have \<open>Im (f x) = 0\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2456
    apply (rule 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
  2457
    using assms
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2458
    by (auto simp add: infsum_Im summable_on_Im 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
  2459
  moreover have \<open>Re (f x) = 0\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2460
    apply (rule 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
  2461
    using assms by (auto simp add: summable_on_Re infsum_Re 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
  2462
  ultimately show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2463
    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
  2464
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2465
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2466
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
  2467
  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
  2468
  assumes "has_sum f A a" and \<open>a \<le> 0\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2469
    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
  2470
  shows "f x = 0"
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2471
  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
  2472
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2473
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
  2474
      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
  2475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2476
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
  2477
  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
  2478
  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
  2479
    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
  2480
  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
  2481
  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
  2482
  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
  2483
  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
  2484
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2485
  have \<open>infsum (\<lambda>x. Re (f x)) A \<le> infsum (\<lambda>x. Re (g x)) B\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2486
    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
  2487
    using assms(3-5) by (auto simp add: summable_on_Re 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
  2488
  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
  2489
    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
  2490
  have \<open>infsum (\<lambda>x. Im (f x)) A = infsum (\<lambda>x. Im (g x)) B\<close>
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2491
    apply (rule infsum_cong_neutral)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2492
    using assms(3-5) by (auto simp add: summable_on_Re 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
  2493
  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
  2494
    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
  2495
  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
  2496
    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
  2497
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2498
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2499
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
  2500
  \<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
  2501
      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
  2502
  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
  2503
  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
  2504
  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
  2505
  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
  2506
  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
  2507
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2508
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2509
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
  2510
  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
  2511
  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
  2512
    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
  2513
  shows "infsum f M \<ge> 0" (is "?lhs \<ge> _")
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2514
  by (metis assms(1) assms(2) infsum_0_simp summable_on_0_simp infsum_mono_complex)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2515
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2516
lemma infsum_cmod:
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2517
  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
  2518
    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
  2519
  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
  2520
proof -
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2521
  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
  2522
  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
  2523
    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
  2524
      using assms summable_on_Re by blast
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2525
    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
  2526
      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
  2527
    finally show \<dots> .
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2528
  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
  2529
  also have \<open>\<dots> = 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
  2530
    apply (rule infsum_cong)
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2531
    using fnn cmod_eq_Re complex_is_Real_iff less_eq_complex_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
  2532
  finally show ?thesis
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2533
    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
  2534
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2535
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2536
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2537
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
  2538
  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
  2539
  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
  2540
proof (rule iffI)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2541
  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
  2542
  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
  2543
    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
  2544
  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
  2545
    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
  2546
  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
  2547
    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
  2548
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2549
  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
  2550
    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
  2551
  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
  2552
    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
  2553
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2554
  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
  2555
    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
  2556
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2557
  have *: \<open>(\<lambda>x. nr x + ni 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
  2558
    apply (rule summable_on_add) by auto
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2559
  show \<open>n summable_on A\<close>
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2560
    apply (rule nonneg_bdd_above_summable_on)
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2561
     apply (simp add: n_def; fail)
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2562
    apply (rule bdd_aboveI[where M=\<open>infsum (\<lambda>x. nr x + ni x) A\<close>])
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2563
    using * n_sum by (auto simp flip: infsum_finite simp: ni_def[abs_def] nr_def[abs_def] intro!: infsum_mono_neutral)
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2564
next
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2565
  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
  2566
    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
  2567
qed
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2568
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2569
lemma summable_countable_complex:
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2570
  fixes f :: \<open>'a \<Rightarrow> complex\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2571
  assumes \<open>f summable_on A\<close>
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2572
  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
  2573
  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
  2574
74475
409ca22dee4c new notion of infinite sums in HOL-Analysis, ordering on complex numbers
eberlm <eberlm@in.tum.de>
parents:
diff changeset
  2575
end
74791
227915e07891 more material for HOL-Analysis.Infinite_Sum
Manuel Eberl <manuel@pruvisto.org>
parents: 74642
diff changeset
  2576