src/HOL/Analysis/Bochner_Integration.thy
author haftmann
Wed, 23 Jun 2021 17:43:31 +0000
changeset 73869 7181130f5872
parent 73536 5131c388a9b0
child 73933 fa92bc604c59
permissions -rw-r--r--
more default simp rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63626
diff changeset
     1
(*  Title:      HOL/Analysis/Bochner_Integration.thy
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
     2
    Author:     Johannes Hölzl, TU München
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
     3
*)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
     4
69722
b5163b2132c5 minor tagging updates in 13 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69700
diff changeset
     5
section \<open>Bochner Integration for Vector-Valued Functions\<close>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
     6
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
     7
theory Bochner_Integration
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
     8
  imports Finite_Product_Measure
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
     9
begin
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    10
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
    11
text \<open>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    12
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    13
In the following development of the Bochner integral we use second countable topologies instead
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    14
of separable spaces. A second countable topology is also separable.
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    15
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
    16
\<close>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    17
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
    18
proposition borel_measurable_implies_sequence_metric:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    19
  fixes f :: "'a \<Rightarrow> 'b :: {metric_space, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    20
  assumes [measurable]: "f \<in> borel_measurable M"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
    21
  shows "\<exists>F. (\<forall>i. simple_function M (F i)) \<and> (\<forall>x\<in>space M. (\<lambda>i. F i x) \<longlonglongrightarrow> f x) \<and>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    22
    (\<forall>i. \<forall>x\<in>space M. dist (F i x) z \<le> 2 * dist (f x) z)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
    23
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    24
  obtain D :: "'b set" where "countable D" and D: "\<And>X. open X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> \<exists>d\<in>D. d \<in> X"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    25
    by (erule countable_dense_setE)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    26
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
    27
  define e where "e = from_nat_into D"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    28
  { fix n x
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    29
    obtain d where "d \<in> D" and d: "d \<in> ball x (1 / Suc n)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    30
      using D[of "ball x (1 / Suc n)"] by auto
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
    31
    from \<open>d \<in> D\<close> D[of UNIV] \<open>countable D\<close> obtain i where "d = e i"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    32
      unfolding e_def by (auto dest: from_nat_into_surj)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    33
    with d have "\<exists>i. dist x (e i) < 1 / Suc n"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    34
      by auto }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    35
  note e = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    36
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
    37
  define A where [abs_def]: "A m n =
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
    38
    {x\<in>space M. dist (f x) (e n) < 1 / (Suc m) \<and> 1 / (Suc m) \<le> dist (f x) z}" for m n
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
    39
  define B where [abs_def]: "B m = disjointed (A m)" for m
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    40
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
    41
  define m where [abs_def]: "m N x = Max {m. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}" for N x
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
    42
  define F where [abs_def]: "F N x =
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
    43
    (if (\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)) \<and> (\<exists>n\<le>N. x \<in> B (m N x) n)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
    44
     then e (LEAST n. x \<in> B (m N x) n) else z)" for N x
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    45
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    46
  have B_imp_A[intro, simp]: "\<And>x m n. x \<in> B m n \<Longrightarrow> x \<in> A m n"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    47
    using disjointed_subset[of "A m" for m] unfolding B_def by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    48
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    49
  { fix m
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    50
    have "\<And>n. A m n \<in> sets M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    51
      by (auto simp: A_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    52
    then have "\<And>n. B m n \<in> sets M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    53
      using sets.range_disjointed_sets[of "A m" M] by (auto simp: B_def) }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    54
  note this[measurable]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    55
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    56
  { fix N i x assume "\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    57
    then have "m N x \<in> {m::nat. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    58
      unfolding m_def by (intro Max_in) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    59
    then have "m N x \<le> N" "\<exists>n\<le>N. x \<in> B (m N x) n"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    60
      by auto }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    61
  note m = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    62
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    63
  { fix j N i x assume "j \<le> N" "i \<le> N" "x \<in> B j i"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    64
    then have "j \<le> m N x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    65
      unfolding m_def by (intro Max_ge) auto }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    66
  note m_upper = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    67
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    68
  show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    69
    unfolding simple_function_def
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    70
  proof (safe intro!: exI[of _ F])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    71
    have [measurable]: "\<And>i. F i \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    72
      unfolding F_def m_def by measurable
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    73
    show "\<And>x i. F i -` {x} \<inter> space M \<in> sets M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    74
      by measurable
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    75
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    76
    { fix i
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    77
      { fix n x assume "x \<in> B (m i x) n"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    78
        then have "(LEAST n. x \<in> B (m i x) n) \<le> n"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    79
          by (intro Least_le)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    80
        also assume "n \<le> i"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    81
        finally have "(LEAST n. x \<in> B (m i x) n) \<le> i" . }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    82
      then have "F i ` space M \<subseteq> {z} \<union> e ` {.. i}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    83
        by (auto simp: F_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    84
      then show "finite (F i ` space M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    85
        by (rule finite_subset) auto }
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    86
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    87
    { fix N i n x assume "i \<le> N" "n \<le> N" "x \<in> B i n"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60066
diff changeset
    88
      then have 1: "\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)" by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    89
      from m[OF this] obtain n where n: "m N x \<le> N" "n \<le> N" "x \<in> B (m N x) n" by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    90
      moreover
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
    91
      define L where "L = (LEAST n. x \<in> B (m N x) n)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    92
      have "dist (f x) (e L) < 1 / Suc (m N x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    93
      proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    94
        have "x \<in> B (m N x) L"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    95
          using n(3) unfolding L_def by (rule LeastI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    96
        then have "x \<in> A (m N x) L"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    97
          by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    98
        then show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
    99
          unfolding A_def by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   100
      qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   101
      ultimately have "dist (f x) (F N x) < 1 / Suc (m N x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   102
        by (auto simp add: F_def L_def) }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   103
    note * = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   104
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   105
    fix x assume "x \<in> space M"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   106
    show "(\<lambda>i. F i x) \<longlonglongrightarrow> f x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   107
    proof cases
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   108
      assume "f x = z"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   109
      then have "\<And>i n. x \<notin> A i n"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   110
        unfolding A_def by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   111
      then have "\<And>i. F i x = z"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   112
        by (auto simp: F_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   113
      then show ?thesis
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
   114
        using \<open>f x = z\<close> by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   115
    next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   116
      assume "f x \<noteq> z"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   117
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   118
      show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   119
      proof (rule tendstoI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   120
        fix e :: real assume "0 < e"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
   121
        with \<open>f x \<noteq> z\<close> obtain n where "1 / Suc n < e" "1 / Suc n < dist (f x) z"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   122
          by (metis dist_nz order_less_trans neq_iff nat_approx_posE)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
   123
        with \<open>x\<in>space M\<close> \<open>f x \<noteq> z\<close> have "x \<in> (\<Union>i. B n i)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   124
          unfolding A_def B_def UN_disjointed_eq using e by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   125
        then obtain i where i: "x \<in> B n i" by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   126
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   127
        show "eventually (\<lambda>i. dist (F i x) (f x) < e) sequentially"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   128
          using eventually_ge_at_top[of "max n i"]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   129
        proof eventually_elim
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   130
          fix j assume j: "max n i \<le> j"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   131
          with i have "dist (f x) (F j x) < 1 / Suc (m j x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   132
            by (intro *[OF _ _ i]) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   133
          also have "\<dots> \<le> 1 / Suc n"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   134
            using j m_upper[OF _ _ i]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   135
            by (auto simp: field_simps)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
   136
          also note \<open>1 / Suc n < e\<close>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   137
          finally show "dist (F j x) (f x) < e"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   138
            by (simp add: less_imp_le dist_commute)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   139
        qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   140
      qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   141
    qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   142
    fix i
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   143
    { fix n m assume "x \<in> A n m"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   144
      then have "dist (e m) (f x) + dist (f x) z \<le> 2 * dist (f x) z"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   145
        unfolding A_def by (auto simp: dist_commute)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   146
      also have "dist (e m) z \<le> dist (e m) (f x) + dist (f x) z"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   147
        by (rule dist_triangle)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   148
      finally (xtrans) have "dist (e m) z \<le> 2 * dist (f x) z" . }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   149
    then show "dist (F i x) z \<le> 2 * dist (f x) z"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   150
      unfolding F_def
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   151
      apply auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   152
      apply (rule LeastI2)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   153
      apply auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   154
      done
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   155
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   156
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   157
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   158
lemma
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   159
  fixes f :: "'a \<Rightarrow> 'b::semiring_1" assumes "finite A"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   160
  shows sum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator (B x) (g x)) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   161
  and sum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator (B x) (g x) * f x) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   162
  unfolding indicator_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   163
  using assms by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   164
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   165
lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   166
  fixes P :: "('a \<Rightarrow> real) \<Rightarrow> bool"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   167
  assumes u: "u \<in> borel_measurable M" "\<And>x. 0 \<le> u x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   168
  assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   169
  assumes mult: "\<And>u c. 0 \<le> c \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   170
  assumes add: "\<And>u v. u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> v x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   171
  assumes seq: "\<And>U. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. 0 \<le> U i x) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. U i x) \<longlonglongrightarrow> u x) \<Longrightarrow> P u"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   172
  shows "P u"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   173
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   174
  have "(\<lambda>x. ennreal (u x)) \<in> borel_measurable M" using u by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   175
  from borel_measurable_implies_simple_function_sequence'[OF this]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   176
  obtain U where U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i x. U i x < top" and
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   177
    sup: "\<And>x. (SUP i. U i x) = ennreal (u x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   178
    by blast
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   179
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   180
  define U' where [abs_def]: "U' i x = indicator (space M) x * enn2real (U i x)" for i x
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   181
  then have U'_sf[measurable]: "\<And>i. simple_function M (U' i)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   182
    using U by (auto intro!: simple_function_compose1[where g=enn2real])
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   183
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   184
  show "P u"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   185
  proof (rule seq)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   186
    show U': "U' i \<in> borel_measurable M" "\<And>x. 0 \<le> U' i x" for i
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   187
      using U by (auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   188
          intro: borel_measurable_simple_function
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   189
          intro!: borel_measurable_enn2real borel_measurable_times
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   190
          simp: U'_def zero_le_mult_iff)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   191
    show "incseq U'"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   192
      using U(2,3)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   193
      by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def enn2real_mono)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   194
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   195
    fix x assume x: "x \<in> space M"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   196
    have "(\<lambda>i. U i x) \<longlonglongrightarrow> (SUP i. U i x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   197
      using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_def le_fun_def)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   198
    moreover have "(\<lambda>i. U i x) = (\<lambda>i. ennreal (U' i x))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   199
      using x U(3) by (auto simp: fun_eq_iff U'_def image_iff eq_commute)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   200
    moreover have "(SUP i. U i x) = ennreal (u x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   201
      using sup u(2) by (simp add: max_def)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   202
    ultimately show "(\<lambda>i. U' i x) \<longlonglongrightarrow> u x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   203
      using u U' by simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   204
  next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   205
    fix i
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   206
    have "U' i ` space M \<subseteq> enn2real ` (U i ` space M)" "finite (U i ` space M)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   207
      unfolding U'_def using U(1) by (auto dest: simple_functionD)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   208
    then have fin: "finite (U' i ` space M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   209
      by (metis finite_subset finite_imageI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   210
    moreover have "\<And>z. {y. U' i z = y \<and> y \<in> U' i ` space M \<and> z \<in> space M} = (if z \<in> space M then {U' i z} else {})"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   211
      by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   212
    ultimately have U': "(\<lambda>z. \<Sum>y\<in>U' i`space M. y * indicator {x\<in>space M. U' i x = y} z) = U' i"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   213
      by (simp add: U'_def fun_eq_iff)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   214
    have "\<And>x. x \<in> U' i ` space M \<Longrightarrow> 0 \<le> x"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   215
      by (auto simp: U'_def)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   216
    with fin have "P (\<lambda>z. \<Sum>y\<in>U' i`space M. y * indicator {x\<in>space M. U' i x = y} z)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   217
    proof induct
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   218
      case empty from set[of "{}"] show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   219
        by (simp add: indicator_def[abs_def])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   220
    next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   221
      case (insert x F)
71840
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71633
diff changeset
   222
      from insert.prems have nonneg: "x \<ge> 0" "\<And>y. y \<in> F \<Longrightarrow> y \<ge> 0"
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71633
diff changeset
   223
        by simp_all
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71633
diff changeset
   224
      hence *: "P (\<lambda>xa. x * indicat_real {x' \<in> space M. U' i x' = x} xa)"
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71633
diff changeset
   225
        by (intro mult set) auto
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71633
diff changeset
   226
      have "P (\<lambda>z. x * indicat_real {x' \<in> space M. U' i x' = x} z + 
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71633
diff changeset
   227
                   (\<Sum>y\<in>F. y * indicat_real {x \<in> space M. U' i x = y} z))"
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71633
diff changeset
   228
        using insert(1-3)
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71633
diff changeset
   229
        by (intro add * sum_nonneg mult_nonneg_nonneg)
73536
5131c388a9b0 simplified definition
haftmann
parents: 73253
diff changeset
   230
           (auto simp: nonneg indicator_def of_bool_def sum_nonneg_eq_0_iff)
71840
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71633
diff changeset
   231
      thus ?case 
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71633
diff changeset
   232
        using insert.hyps by (subst sum.insert) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   233
    qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   234
    with U' show "P (U' i)" by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   235
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   236
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   237
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   238
lemma scaleR_cong_right:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   239
  fixes x :: "'a :: real_vector"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   240
  shows "(x \<noteq> 0 \<Longrightarrow> r = p) \<Longrightarrow> r *\<^sub>R x = p *\<^sub>R x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   241
  by (cases "x = 0") auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   242
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   243
inductive simple_bochner_integrable :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> bool" for M f where
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   244
  "simple_function M f \<Longrightarrow> emeasure M {y\<in>space M. f y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   245
    simple_bochner_integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   246
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   247
lemma simple_bochner_integrable_compose2:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   248
  assumes p_0: "p 0 0 = 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   249
  shows "simple_bochner_integrable M f \<Longrightarrow> simple_bochner_integrable M g \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   250
    simple_bochner_integrable M (\<lambda>x. p (f x) (g x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   251
proof (safe intro!: simple_bochner_integrable.intros elim!: simple_bochner_integrable.cases del: notI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   252
  assume sf: "simple_function M f" "simple_function M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   253
  then show "simple_function M (\<lambda>x. p (f x) (g x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   254
    by (rule simple_function_compose2)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   255
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   256
  from sf have [measurable]:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   257
      "f \<in> measurable M (count_space UNIV)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   258
      "g \<in> measurable M (count_space UNIV)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   259
    by (auto intro: measurable_simple_function)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   260
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   261
  assume fin: "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" "emeasure M {y \<in> space M. g y \<noteq> 0} \<noteq> \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   262
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   263
  have "emeasure M {x\<in>space M. p (f x) (g x) \<noteq> 0} \<le>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   264
      emeasure M ({x\<in>space M. f x \<noteq> 0} \<union> {x\<in>space M. g x \<noteq> 0})"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   265
    by (intro emeasure_mono) (auto simp: p_0)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   266
  also have "\<dots> \<le> emeasure M {x\<in>space M. f x \<noteq> 0} + emeasure M {x\<in>space M. g x \<noteq> 0}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   267
    by (intro emeasure_subadditive) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   268
  finally show "emeasure M {y \<in> space M. p (f y) (g y) \<noteq> 0} \<noteq> \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   269
    using fin by (auto simp: top_unique)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   270
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   271
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   272
lemma simple_function_finite_support:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   273
  assumes f: "simple_function M f" and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" and nn: "\<And>x. 0 \<le> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   274
  shows "emeasure M {x\<in>space M. f x \<noteq> 0} \<noteq> \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   275
proof cases
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   276
  from f have meas[measurable]: "f \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   277
    by (rule borel_measurable_simple_function)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   278
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   279
  assume non_empty: "\<exists>x\<in>space M. f x \<noteq> 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   280
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   281
  define m where "m = Min (f`space M - {0})"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   282
  have "m \<in> f`space M - {0}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   283
    unfolding m_def using f non_empty by (intro Min_in) (auto simp: simple_function_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   284
  then have m: "0 < m"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   285
    using nn by (auto simp: less_le)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   286
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   287
  from m have "m * emeasure M {x\<in>space M. 0 \<noteq> f x} =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   288
    (\<integral>\<^sup>+x. m * indicator {x\<in>space M. 0 \<noteq> f x} x \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   289
    using f by (intro nn_integral_cmult_indicator[symmetric]) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   290
  also have "\<dots> \<le> (\<integral>\<^sup>+x. f x \<partial>M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   291
    using AE_space
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   292
  proof (intro nn_integral_mono_AE, eventually_elim)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   293
    fix x assume "x \<in> space M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   294
    with nn show "m * indicator {x \<in> space M. 0 \<noteq> f x} x \<le> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   295
      using f by (auto split: split_indicator simp: simple_function_def m_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   296
  qed
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
   297
  also note \<open>\<dots> < \<infinity>\<close>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   298
  finally show ?thesis
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   299
    using m by (auto simp: ennreal_mult_less_top)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   300
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   301
  assume "\<not> (\<exists>x\<in>space M. f x \<noteq> 0)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   302
  with nn have *: "{x\<in>space M. f x \<noteq> 0} = {}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   303
    by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   304
  show ?thesis unfolding * by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   305
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   306
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   307
lemma simple_bochner_integrableI_bounded:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   308
  assumes f: "simple_function M f" and fin: "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   309
  shows "simple_bochner_integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   310
proof
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   311
  have "emeasure M {y \<in> space M. ennreal (norm (f y)) \<noteq> 0} \<noteq> \<infinity>"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   312
  proof (rule simple_function_finite_support)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   313
    show "simple_function M (\<lambda>x. ennreal (norm (f x)))"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   314
      using f by (rule simple_function_compose1)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   315
    show "(\<integral>\<^sup>+ y. ennreal (norm (f y)) \<partial>M) < \<infinity>" by fact
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   316
  qed simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   317
  then show "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   318
qed fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   319
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69739
diff changeset
   320
definition\<^marker>\<open>tag important\<close> simple_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> 'b" where
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   321
  "simple_bochner_integral M f = (\<Sum>y\<in>f`space M. measure M {x\<in>space M. f x = y} *\<^sub>R y)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   322
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   323
proposition simple_bochner_integral_partition:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   324
  assumes f: "simple_bochner_integrable M f" and g: "simple_function M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   325
  assumes sub: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow> g x = g y \<Longrightarrow> f x = f y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   326
  assumes v: "\<And>x. x \<in> space M \<Longrightarrow> f x = v (g x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   327
  shows "simple_bochner_integral M f = (\<Sum>y\<in>g ` space M. measure M {x\<in>space M. g x = y} *\<^sub>R v y)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   328
    (is "_ = ?r")
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   329
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   330
  from f g have [simp]: "finite (f`space M)" "finite (g`space M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   331
    by (auto simp: simple_function_def elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   332
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   333
  from f have [measurable]: "f \<in> measurable M (count_space UNIV)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   334
    by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   335
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   336
  from g have [measurable]: "g \<in> measurable M (count_space UNIV)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   337
    by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   338
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   339
  { fix y assume "y \<in> space M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   340
    then have "f ` space M \<inter> {i. \<exists>x\<in>space M. i = f x \<and> g y = g x} = {v (g y)}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   341
      by (auto cong: sub simp: v[symmetric]) }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   342
  note eq = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   343
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   344
  have "simple_bochner_integral M f =
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   345
    (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   346
      if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} else 0) *\<^sub>R y)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   347
    unfolding simple_bochner_integral_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   348
  proof (safe intro!: sum.cong scaleR_cong_right)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   349
    fix y assume y: "y \<in> space M" "f y \<noteq> 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   350
    have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   351
        {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   352
      by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   353
    have eq:"{x \<in> space M. f x = f y} =
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   354
        (\<Union>i\<in>{z. \<exists>x\<in>space M. f y = f x \<and> z = g x}. {x \<in> space M. g x = i})"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   355
      by (auto simp: eq_commute cong: sub rev_conj_cong)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   356
    have "finite (g`space M)" by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   357
    then have "finite {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   358
      by (rule rev_finite_subset) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   359
    moreover
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   360
    { fix x assume "x \<in> space M" "f x = f y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   361
      then have "x \<in> space M" "f x \<noteq> 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   362
        using y by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   363
      then have "emeasure M {y \<in> space M. g y = g x} \<le> emeasure M {y \<in> space M. f y \<noteq> 0}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   364
        by (auto intro!: emeasure_mono cong: sub)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   365
      then have "emeasure M {xa \<in> space M. g xa = g x} < \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   366
        using f by (auto simp: simple_bochner_integrable.simps less_top) }
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   367
    ultimately
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   368
    show "measure M {x \<in> space M. f x = f y} =
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   369
      (\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then measure M {x \<in> space M. g x = z} else 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   370
      apply (simp add: sum.If_cases eq)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   371
      apply (subst measure_finite_Union[symmetric])
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   372
      apply (auto simp: disjoint_family_on_def less_top)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   373
      done
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   374
  qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   375
  also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   376
      if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} *\<^sub>R y else 0))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   377
    by (auto intro!: sum.cong simp: scaleR_sum_left)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   378
  also have "\<dots> = ?r"
66804
3f9bb52082c4 avoid name clashes on interpretation of abstract locales
haftmann
parents: 66447
diff changeset
   379
    by (subst sum.swap)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   380
       (auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   381
  finally show "simple_bochner_integral M f = ?r" .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   382
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   383
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   384
lemma simple_bochner_integral_add:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   385
  assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   386
  shows "simple_bochner_integral M (\<lambda>x. f x + g x) =
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   387
    simple_bochner_integral M f + simple_bochner_integral M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   388
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   389
  from f g have "simple_bochner_integral M (\<lambda>x. f x + g x) =
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   390
    (\<Sum>y\<in>(\<lambda>x. (f x, g x)) ` space M. measure M {x \<in> space M. (f x, g x) = y} *\<^sub>R (fst y + snd y))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   391
    by (intro simple_bochner_integral_partition)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   392
       (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   393
  moreover from f g have "simple_bochner_integral M f =
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   394
    (\<Sum>y\<in>(\<lambda>x. (f x, g x)) ` space M. measure M {x \<in> space M. (f x, g x) = y} *\<^sub>R fst y)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   395
    by (intro simple_bochner_integral_partition)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   396
       (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   397
  moreover from f g have "simple_bochner_integral M g =
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   398
    (\<Sum>y\<in>(\<lambda>x. (f x, g x)) ` space M. measure M {x \<in> space M. (f x, g x) = y} *\<^sub>R snd y)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   399
    by (intro simple_bochner_integral_partition)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   400
       (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   401
  ultimately show ?thesis
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   402
    by (simp add: sum.distrib[symmetric] scaleR_add_right)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   403
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   404
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   405
lemma simple_bochner_integral_linear:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67399
diff changeset
   406
  assumes "linear f"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   407
  assumes g: "simple_bochner_integrable M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   408
  shows "simple_bochner_integral M (\<lambda>x. f (g x)) = f (simple_bochner_integral M g)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   409
proof -
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67399
diff changeset
   410
  interpret linear f by fact
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   411
  from g have "simple_bochner_integral M (\<lambda>x. f (g x)) =
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   412
    (\<Sum>y\<in>g ` space M. measure M {x \<in> space M. g x = y} *\<^sub>R f y)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   413
    by (intro simple_bochner_integral_partition)
71633
07bec530f02e cleaned proofs
nipkow
parents: 70532
diff changeset
   414
       (auto simp: simple_bochner_integrable_compose2[where p="\<lambda>x y. f x"]
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   415
             elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   416
  also have "\<dots> = f (simple_bochner_integral M g)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67399
diff changeset
   417
    by (simp add: simple_bochner_integral_def sum scale)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   418
  finally show ?thesis .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   419
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   420
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   421
lemma simple_bochner_integral_minus:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   422
  assumes f: "simple_bochner_integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   423
  shows "simple_bochner_integral M (\<lambda>x. - f x) = - simple_bochner_integral M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   424
proof -
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67399
diff changeset
   425
  from linear_uminus f show ?thesis
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   426
    by (rule simple_bochner_integral_linear)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   427
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   428
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   429
lemma simple_bochner_integral_diff:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   430
  assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   431
  shows "simple_bochner_integral M (\<lambda>x. f x - g x) =
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   432
    simple_bochner_integral M f - simple_bochner_integral M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   433
  unfolding diff_conv_add_uminus using f g
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   434
  by (subst simple_bochner_integral_add)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   435
     (auto simp: simple_bochner_integral_minus simple_bochner_integrable_compose2[where p="\<lambda>x y. - y"])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   436
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   437
lemma simple_bochner_integral_norm_bound:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   438
  assumes f: "simple_bochner_integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   439
  shows "norm (simple_bochner_integral M f) \<le> simple_bochner_integral M (\<lambda>x. norm (f x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   440
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   441
  have "norm (simple_bochner_integral M f) \<le>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   442
    (\<Sum>y\<in>f ` space M. norm (measure M {x \<in> space M. f x = y} *\<^sub>R y))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   443
    unfolding simple_bochner_integral_def by (rule norm_sum)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   444
  also have "\<dots> = (\<Sum>y\<in>f ` space M. measure M {x \<in> space M. f x = y} *\<^sub>R norm y)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   445
    by simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   446
  also have "\<dots> = simple_bochner_integral M (\<lambda>x. norm (f x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   447
    using f
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   448
    by (intro simple_bochner_integral_partition[symmetric])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   449
       (auto intro: f simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   450
  finally show ?thesis .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   451
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   452
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   453
lemma simple_bochner_integral_nonneg[simp]:
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   454
  fixes f :: "'a \<Rightarrow> real"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   455
  shows "(\<And>x. 0 \<le> f x) \<Longrightarrow> 0 \<le> simple_bochner_integral M f"
65680
378a2f11bec9 Simplification of some proofs. Also key lemmas using !! rather than ! in premises
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
   456
  by (force simp add: simple_bochner_integral_def intro: sum_nonneg)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   457
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   458
lemma simple_bochner_integral_eq_nn_integral:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   459
  assumes f: "simple_bochner_integrable M f" "\<And>x. 0 \<le> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   460
  shows "simple_bochner_integral M f = (\<integral>\<^sup>+x. f x \<partial>M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   461
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   462
  { fix x y z have "(x \<noteq> 0 \<Longrightarrow> y = z) \<Longrightarrow> ennreal x * y = ennreal x * z"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   463
      by (cases "x = 0") (auto simp: zero_ennreal_def[symmetric]) }
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   464
  note ennreal_cong_mult = this
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   465
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   466
  have [measurable]: "f \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   467
    using f(1) by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   468
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   469
  { fix y assume y: "y \<in> space M" "f y \<noteq> 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   470
    have "ennreal (measure M {x \<in> space M. f x = f y}) = emeasure M {x \<in> space M. f x = f y}"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   471
    proof (rule emeasure_eq_ennreal_measure[symmetric])
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   472
      have "emeasure M {x \<in> space M. f x = f y} \<le> emeasure M {x \<in> space M. f x \<noteq> 0}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   473
        using y by (intro emeasure_mono) auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   474
      with f show "emeasure M {x \<in> space M. f x = f y} \<noteq> top"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   475
        by (auto simp: simple_bochner_integrable.simps top_unique)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   476
    qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   477
    moreover have "{x \<in> space M. f x = f y} = (\<lambda>x. ennreal (f x)) -` {ennreal (f y)} \<inter> space M"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   478
      using f by auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   479
    ultimately have "ennreal (measure M {x \<in> space M. f x = f y}) =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   480
          emeasure M ((\<lambda>x. ennreal (f x)) -` {ennreal (f y)} \<inter> space M)" by simp }
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   481
  with f have "simple_bochner_integral M f = (\<integral>\<^sup>Sx. f x \<partial>M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   482
    unfolding simple_integral_def
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   483
    by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ennreal (f x)" and v=enn2real])
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   484
       (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   485
             intro!: sum.cong ennreal_cong_mult
68046
6aba668aea78 new simp modifier: reorient
nipkow
parents: 67399
diff changeset
   486
             simp: ac_simps ennreal_mult
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68073
diff changeset
   487
             simp flip: sum_ennreal)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   488
  also have "\<dots> = (\<integral>\<^sup>+x. f x \<partial>M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   489
    using f
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   490
    by (intro nn_integral_eq_simple_integral[symmetric])
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   491
       (auto simp: simple_function_compose1 simple_bochner_integrable.simps)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   492
  finally show ?thesis .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   493
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   494
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   495
lemma simple_bochner_integral_bounded:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   496
  fixes f :: "'a \<Rightarrow> 'b::{real_normed_vector, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   497
  assumes f[measurable]: "f \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   498
  assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   499
  shows "ennreal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \<le>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   500
    (\<integral>\<^sup>+ x. norm (f x - s x) \<partial>M) + (\<integral>\<^sup>+ x. norm (f x - t x) \<partial>M)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   501
    (is "ennreal (norm (?s - ?t)) \<le> ?S + ?T")
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   502
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   503
  have [measurable]: "s \<in> borel_measurable M" "t \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   504
    using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   505
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   506
  have "ennreal (norm (?s - ?t)) = norm (simple_bochner_integral M (\<lambda>x. s x - t x))"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   507
    using s t by (subst simple_bochner_integral_diff) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   508
  also have "\<dots> \<le> simple_bochner_integral M (\<lambda>x. norm (s x - t x))"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66804
diff changeset
   509
    using simple_bochner_integrable_compose2[of "(-)" M "s" "t"] s t
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   510
    by (auto intro!: simple_bochner_integral_norm_bound)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   511
  also have "\<dots> = (\<integral>\<^sup>+x. norm (s x - t x) \<partial>M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   512
    using simple_bochner_integrable_compose2[of "\<lambda>x y. norm (x - y)" M "s" "t"] s t
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   513
    by (auto intro!: simple_bochner_integral_eq_nn_integral)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   514
  also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \<partial>M)"
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68073
diff changeset
   515
    by (auto intro!: nn_integral_mono simp flip: ennreal_plus)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   516
       (metis (erased, hide_lams) add_diff_cancel_left add_diff_eq diff_add_eq order_trans
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   517
              norm_minus_commute norm_triangle_ineq4 order_refl)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   518
  also have "\<dots> = ?S + ?T"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   519
   by (rule nn_integral_add) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   520
  finally show ?thesis .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   521
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   522
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   523
inductive has_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b::{real_normed_vector, second_countable_topology} \<Rightarrow> bool"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   524
  for M f x where
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   525
  "f \<in> borel_measurable M \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   526
    (\<And>i. simple_bochner_integrable M (s i)) \<Longrightarrow>
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   527
    (\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0 \<Longrightarrow>
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   528
    (\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x \<Longrightarrow>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   529
    has_bochner_integral M f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   530
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   531
lemma has_bochner_integral_cong:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   532
  assumes "M = N" "\<And>x. x \<in> space N \<Longrightarrow> f x = g x" "x = y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   533
  shows "has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral N g y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   534
  unfolding has_bochner_integral.simps assms(1,3)
69546
27dae626822b prefer naming convention from datatype package for strong congruence rules
haftmann
parents: 69144
diff changeset
   535
  using assms(2) by (simp cong: measurable_cong_simp nn_integral_cong_simp)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   536
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   537
lemma has_bochner_integral_cong_AE:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   538
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   539
    has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral M g x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   540
  unfolding has_bochner_integral.simps
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   541
  by (intro arg_cong[where f=Ex] ext conj_cong rev_conj_cong refl arg_cong[where f="\<lambda>x. x \<longlonglongrightarrow> 0"]
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   542
            nn_integral_cong_AE)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   543
     auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   544
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   545
lemma borel_measurable_has_bochner_integral:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   546
  "has_bochner_integral M f x \<Longrightarrow> f \<in> borel_measurable M"
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   547
  by (rule has_bochner_integral.cases)
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   548
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   549
lemma borel_measurable_has_bochner_integral'[measurable_dest]:
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   550
  "has_bochner_integral M f x \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N"
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   551
  using borel_measurable_has_bochner_integral[measurable] by measurable
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   552
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   553
lemma has_bochner_integral_simple_bochner_integrable:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   554
  "simple_bochner_integrable M f \<Longrightarrow> has_bochner_integral M f (simple_bochner_integral M f)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   555
  by (rule has_bochner_integral.intros[where s="\<lambda>_. f"])
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   556
     (auto intro: borel_measurable_simple_function
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   557
           elim: simple_bochner_integrable.cases
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   558
           simp: zero_ennreal_def[symmetric])
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   559
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   560
lemma has_bochner_integral_real_indicator:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   561
  assumes [measurable]: "A \<in> sets M" and A: "emeasure M A < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   562
  shows "has_bochner_integral M (indicator A) (measure M A)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   563
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   564
  have sbi: "simple_bochner_integrable M (indicator A::'a \<Rightarrow> real)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   565
  proof
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   566
    have "{y \<in> space M. (indicator A y::real) \<noteq> 0} = A"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
   567
      using sets.sets_into_space[OF \<open>A\<in>sets M\<close>] by (auto split: split_indicator)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   568
    then show "emeasure M {y \<in> space M. (indicator A y::real) \<noteq> 0} \<noteq> \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   569
      using A by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   570
  qed (rule simple_function_indicator assms)+
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   571
  moreover have "simple_bochner_integral M (indicator A) = measure M A"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   572
    using simple_bochner_integral_eq_nn_integral[OF sbi] A
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   573
    by (simp add: ennreal_indicator emeasure_eq_ennreal_measure)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   574
  ultimately show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   575
    by (metis has_bochner_integral_simple_bochner_integrable)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   576
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   577
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   578
lemma has_bochner_integral_add[intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   579
  "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M g y \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   580
    has_bochner_integral M (\<lambda>x. f x + g x) (x + y)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   581
proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   582
  fix sf sg
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   583
  assume f_sf: "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - sf i x) \<partial>M) \<longlonglongrightarrow> 0"
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   584
  assume g_sg: "(\<lambda>i. \<integral>\<^sup>+ x. norm (g x - sg i x) \<partial>M) \<longlonglongrightarrow> 0"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   585
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   586
  assume sf: "\<forall>i. simple_bochner_integrable M (sf i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   587
    and sg: "\<forall>i. simple_bochner_integrable M (sg i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   588
  then have [measurable]: "\<And>i. sf i \<in> borel_measurable M" "\<And>i. sg i \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   589
    by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   590
  assume [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   591
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   592
  show "\<And>i. simple_bochner_integrable M (\<lambda>x. sf i x + sg i x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   593
    using sf sg by (simp add: simple_bochner_integrable_compose2)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   594
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   595
  show "(\<lambda>i. \<integral>\<^sup>+ x. (norm (f x + g x - (sf i x + sg i x))) \<partial>M) \<longlonglongrightarrow> 0"
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   596
    (is "?f \<longlonglongrightarrow> 0")
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   597
  proof (rule tendsto_sandwich)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   598
    show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   599
      by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   600
    show "eventually (\<lambda>i. ?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) \<partial>M) + \<integral>\<^sup>+ x. (norm (g x - sg i x)) \<partial>M) sequentially"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   601
      (is "eventually (\<lambda>i. ?f i \<le> ?g i) sequentially")
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   602
    proof (intro always_eventually allI)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   603
      fix i have "?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   604
        by (auto intro!: nn_integral_mono norm_diff_triangle_ineq
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68073
diff changeset
   605
                 simp flip: ennreal_plus)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   606
      also have "\<dots> = ?g i"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   607
        by (intro nn_integral_add) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   608
      finally show "?f i \<le> ?g i" .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   609
    qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   610
    show "?g \<longlonglongrightarrow> 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   611
      using tendsto_add[OF f_sf g_sg] by simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   612
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   613
qed (auto simp: simple_bochner_integral_add tendsto_add)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   614
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   615
lemma has_bochner_integral_bounded_linear:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   616
  assumes "bounded_linear T"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   617
  shows "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M (\<lambda>x. T (f x)) (T x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   618
proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   619
  interpret T: bounded_linear T by fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   620
  have [measurable]: "T \<in> borel_measurable borel"
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
   621
    by (intro borel_measurable_continuous_onI T.continuous_on continuous_on_id)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   622
  assume [measurable]: "f \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   623
  then show "(\<lambda>x. T (f x)) \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   624
    by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   625
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   626
  fix s assume f_s: "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   627
  assume s: "\<forall>i. simple_bochner_integrable M (s i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   628
  then show "\<And>i. simple_bochner_integrable M (\<lambda>x. T (s i x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   629
    by (auto intro: simple_bochner_integrable_compose2 T.zero)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   630
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   631
  have [measurable]: "\<And>i. s i \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   632
    using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   633
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   634
  obtain K where K: "K > 0" "\<And>x i. norm (T (f x) - T (s i x)) \<le> norm (f x - s i x) * K"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   635
    using T.pos_bounded by (auto simp: T.diff[symmetric])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   636
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   637
  show "(\<lambda>i. \<integral>\<^sup>+ x. norm (T (f x) - T (s i x)) \<partial>M) \<longlonglongrightarrow> 0"
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   638
    (is "?f \<longlonglongrightarrow> 0")
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   639
  proof (rule tendsto_sandwich)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   640
    show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   641
      by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   642
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   643
    show "eventually (\<lambda>i. ?f i \<le> K * (\<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M)) sequentially"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   644
      (is "eventually (\<lambda>i. ?f i \<le> ?g i) sequentially")
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   645
    proof (intro always_eventually allI)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   646
      fix i have "?f i \<le> (\<integral>\<^sup>+ x. ennreal K * norm (f x - s i x) \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   647
        using K by (intro nn_integral_mono) (auto simp: ac_simps ennreal_mult[symmetric])
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   648
      also have "\<dots> = ?g i"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   649
        using K by (intro nn_integral_cmult) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   650
      finally show "?f i \<le> ?g i" .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   651
    qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   652
    show "?g \<longlonglongrightarrow> 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   653
      using ennreal_tendsto_cmult[OF _ f_s] by simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   654
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   655
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   656
  assume "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x"
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   657
  with s show "(\<lambda>i. simple_bochner_integral M (\<lambda>x. T (s i x))) \<longlonglongrightarrow> T x"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67399
diff changeset
   658
    by (auto intro!: T.tendsto simp: simple_bochner_integral_linear T.linear_axioms)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   659
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   660
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   661
lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\<lambda>x. 0) 0"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   662
  by (auto intro!: has_bochner_integral.intros[where s="\<lambda>_ _. 0"]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   663
           simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   664
                 simple_bochner_integral_def image_constant_conv)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   665
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   666
lemma has_bochner_integral_scaleR_left[intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   667
  "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x *\<^sub>R c) (x *\<^sub>R c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   668
  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   669
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   670
lemma has_bochner_integral_scaleR_right[intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   671
  "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c *\<^sub>R f x) (c *\<^sub>R x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   672
  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   673
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   674
lemma has_bochner_integral_mult_left[intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   675
  fixes c :: "_::{real_normed_algebra,second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   676
  shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x * c) (x * c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   677
  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   678
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   679
lemma has_bochner_integral_mult_right[intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   680
  fixes c :: "_::{real_normed_algebra,second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   681
  shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c * f x) (c * x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   682
  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   683
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   684
lemmas has_bochner_integral_divide =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   685
  has_bochner_integral_bounded_linear[OF bounded_linear_divide]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   686
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   687
lemma has_bochner_integral_divide_zero[intro]:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59587
diff changeset
   688
  fixes c :: "_::{real_normed_field, field, second_countable_topology}"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   689
  shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x / c) (x / c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   690
  using has_bochner_integral_divide by (cases "c = 0") auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   691
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   692
lemma has_bochner_integral_inner_left[intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   693
  "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x \<bullet> c) (x \<bullet> c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   694
  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   695
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   696
lemma has_bochner_integral_inner_right[intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   697
  "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c \<bullet> f x) (c \<bullet> x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   698
  by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   699
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   700
lemmas has_bochner_integral_minus =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   701
  has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   702
lemmas has_bochner_integral_Re =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   703
  has_bochner_integral_bounded_linear[OF bounded_linear_Re]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   704
lemmas has_bochner_integral_Im =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   705
  has_bochner_integral_bounded_linear[OF bounded_linear_Im]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   706
lemmas has_bochner_integral_cnj =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   707
  has_bochner_integral_bounded_linear[OF bounded_linear_cnj]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   708
lemmas has_bochner_integral_of_real =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   709
  has_bochner_integral_bounded_linear[OF bounded_linear_of_real]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   710
lemmas has_bochner_integral_fst =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   711
  has_bochner_integral_bounded_linear[OF bounded_linear_fst]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   712
lemmas has_bochner_integral_snd =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   713
  has_bochner_integral_bounded_linear[OF bounded_linear_snd]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   714
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   715
lemma has_bochner_integral_indicator:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   716
  "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   717
    has_bochner_integral M (\<lambda>x. indicator A x *\<^sub>R c) (measure M A *\<^sub>R c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   718
  by (intro has_bochner_integral_scaleR_left has_bochner_integral_real_indicator)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   719
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   720
lemma has_bochner_integral_diff:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   721
  "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M g y \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   722
    has_bochner_integral M (\<lambda>x. f x - g x) (x - y)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   723
  unfolding diff_conv_add_uminus
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   724
  by (intro has_bochner_integral_add has_bochner_integral_minus)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   725
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   726
lemma has_bochner_integral_sum:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   727
  "(\<And>i. i \<in> I \<Longrightarrow> has_bochner_integral M (f i) (x i)) \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   728
    has_bochner_integral M (\<lambda>x. \<Sum>i\<in>I. f i x) (\<Sum>i\<in>I. x i)"
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
   729
  by (induct I rule: infinite_finite_induct) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   730
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   731
proposition has_bochner_integral_implies_finite_norm:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   732
  "has_bochner_integral M f x \<Longrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   733
proof (elim has_bochner_integral.cases)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   734
  fix s v
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   735
  assume [measurable]: "f \<in> borel_measurable M" and s: "\<And>i. simple_bochner_integrable M (s i)" and
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   736
    lim_0: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   737
  from order_tendstoD[OF lim_0, of "\<infinity>"]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   738
  obtain i where f_s_fin: "(\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) < \<infinity>"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   739
    by (auto simp: eventually_sequentially)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   740
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   741
  have [measurable]: "\<And>i. s i \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   742
    using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   743
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   744
  define m where "m = (if space M = {} then 0 else Max ((\<lambda>x. norm (s i x))`space M))"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   745
  have "finite (s i ` space M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   746
    using s by (auto simp: simple_function_def simple_bochner_integrable.simps)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   747
  then have "finite (norm ` s i ` space M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   748
    by (rule finite_imageI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   749
  then have "\<And>x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> m" "0 \<le> m"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   750
    by (auto simp: m_def image_comp comp_def Max_ge_iff)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   751
  then have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal m * indicator {x\<in>space M. s i x \<noteq> 0} x \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   752
    by (auto split: split_indicator intro!: Max_ge nn_integral_mono simp:)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   753
  also have "\<dots> < \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   754
    using s by (subst nn_integral_cmult_indicator) (auto simp: \<open>0 \<le> m\<close> simple_bochner_integrable.simps ennreal_mult_less_top less_top)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   755
  finally have s_fin: "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>" .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   756
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   757
  have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \<partial>M)"
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68073
diff changeset
   758
    by (auto intro!: nn_integral_mono simp flip: ennreal_plus)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   759
       (metis add.commute norm_triangle_sub)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   760
  also have "\<dots> = (\<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) + (\<integral>\<^sup>+x. norm (s i x) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   761
    by (rule nn_integral_add) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   762
  also have "\<dots> < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   763
    using s_fin f_s_fin by auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   764
  finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   765
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   766
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   767
proposition has_bochner_integral_norm_bound:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   768
  assumes i: "has_bochner_integral M f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   769
  shows "norm x \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   770
using assms proof
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   771
  fix s assume
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   772
    x: "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x" (is "?s \<longlonglongrightarrow> x") and
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   773
    s[simp]: "\<And>i. simple_bochner_integrable M (s i)" and
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   774
    lim: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" and
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   775
    f[measurable]: "f \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   776
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   777
  have [measurable]: "\<And>i. s i \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   778
    using s by (auto simp: simple_bochner_integrable.simps intro: borel_measurable_simple_function)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   779
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   780
  show "norm x \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   781
  proof (rule LIMSEQ_le)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   782
    show "(\<lambda>i. ennreal (norm (?s i))) \<longlonglongrightarrow> norm x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   783
      using x by (auto simp: tendsto_ennreal_iff intro: tendsto_intros)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   784
    show "\<exists>N. \<forall>n\<ge>N. norm (?s n) \<le> (\<integral>\<^sup>+x. norm (f x - s n x) \<partial>M) + (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   785
      (is "\<exists>N. \<forall>n\<ge>N. _ \<le> ?t n")
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   786
    proof (intro exI allI impI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   787
      fix n
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   788
      have "ennreal (norm (?s n)) \<le> simple_bochner_integral M (\<lambda>x. norm (s n x))"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   789
        by (auto intro!: simple_bochner_integral_norm_bound)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   790
      also have "\<dots> = (\<integral>\<^sup>+x. norm (s n x) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   791
        by (intro simple_bochner_integral_eq_nn_integral)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   792
           (auto intro: s simple_bochner_integrable_compose2)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   793
      also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \<partial>M)"
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68073
diff changeset
   794
        by (auto intro!: nn_integral_mono simp flip: ennreal_plus)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57447
diff changeset
   795
           (metis add.commute norm_minus_commute norm_triangle_sub)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   796
      also have "\<dots> = ?t n"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   797
        by (rule nn_integral_add) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   798
      finally show "norm (?s n) \<le> ?t n" .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   799
    qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   800
    have "?t \<longlonglongrightarrow> 0 + (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   801
      using has_bochner_integral_implies_finite_norm[OF i]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   802
      by (intro tendsto_add tendsto_const lim)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   803
    then show "?t \<longlonglongrightarrow> \<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   804
      by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   805
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   806
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   807
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   808
lemma has_bochner_integral_eq:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   809
  "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M f y \<Longrightarrow> x = y"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   810
proof (elim has_bochner_integral.cases)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   811
  assume f[measurable]: "f \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   812
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   813
  fix s t
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   814
  assume "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?S \<longlonglongrightarrow> 0")
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   815
  assume "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - t i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?T \<longlonglongrightarrow> 0")
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   816
  assume s: "\<And>i. simple_bochner_integrable M (s i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   817
  assume t: "\<And>i. simple_bochner_integrable M (t i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   818
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   819
  have [measurable]: "\<And>i. s i \<in> borel_measurable M" "\<And>i. t i \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   820
    using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   821
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   822
  let ?s = "\<lambda>i. simple_bochner_integral M (s i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   823
  let ?t = "\<lambda>i. simple_bochner_integral M (t i)"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   824
  assume "?s \<longlonglongrightarrow> x" "?t \<longlonglongrightarrow> y"
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   825
  then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> norm (x - y)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   826
    by (intro tendsto_intros)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   827
  moreover
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   828
  have "(\<lambda>i. ennreal (norm (?s i - ?t i))) \<longlonglongrightarrow> ennreal 0"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   829
  proof (rule tendsto_sandwich)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   830
    show "eventually (\<lambda>i. 0 \<le> ennreal (norm (?s i - ?t i))) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> ennreal 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   831
      by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   832
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   833
    show "eventually (\<lambda>i. norm (?s i - ?t i) \<le> ?S i + ?T i) sequentially"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   834
      by (intro always_eventually allI simple_bochner_integral_bounded s t f)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   835
    show "(\<lambda>i. ?S i + ?T i) \<longlonglongrightarrow> ennreal 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   836
      using tendsto_add[OF \<open>?S \<longlonglongrightarrow> 0\<close> \<open>?T \<longlonglongrightarrow> 0\<close>] by simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   837
  qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   838
  then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> 0"
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68073
diff changeset
   839
    by (simp flip: ennreal_0)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   840
  ultimately have "norm (x - y) = 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   841
    by (rule LIMSEQ_unique)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   842
  then show "x = y" by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   843
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   844
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   845
lemma has_bochner_integralI_AE:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   846
  assumes f: "has_bochner_integral M f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   847
    and g: "g \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   848
    and ae: "AE x in M. f x = g x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   849
  shows "has_bochner_integral M g x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   850
  using f
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   851
proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   852
  fix s assume "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   853
  also have "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) = (\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   854
    using ae
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   855
    by (intro ext nn_integral_cong_AE, eventually_elim) simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   856
  finally show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" .
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   857
qed (auto intro: g)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   858
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   859
lemma has_bochner_integral_eq_AE:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   860
  assumes f: "has_bochner_integral M f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   861
    and g: "has_bochner_integral M g y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   862
    and ae: "AE x in M. f x = g x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   863
  shows "x = y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   864
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   865
  from assms have "has_bochner_integral M g x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   866
    by (auto intro: has_bochner_integralI_AE)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   867
  from this g show "x = y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   868
    by (rule has_bochner_integral_eq)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   869
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   870
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   871
lemma simple_bochner_integrable_restrict_space:
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   872
  fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   873
  assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   874
  shows "simple_bochner_integrable (restrict_space M \<Omega>) f \<longleftrightarrow>
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   875
    simple_bochner_integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   876
  by (simp add: simple_bochner_integrable.simps space_restrict_space
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   877
    simple_function_restrict_space[OF \<Omega>] emeasure_restrict_space[OF \<Omega>] Collect_restrict
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   878
    indicator_eq_0_iff conj_left_commute)
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   879
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   880
lemma simple_bochner_integral_restrict_space:
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   881
  fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   882
  assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   883
  assumes f: "simple_bochner_integrable (restrict_space M \<Omega>) f"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   884
  shows "simple_bochner_integral (restrict_space M \<Omega>) f =
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   885
    simple_bochner_integral M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   886
proof -
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   887
  have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x)`space M)"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   888
    using f simple_bochner_integrable_restrict_space[OF \<Omega>, of f]
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   889
    by (simp add: simple_bochner_integrable.simps simple_function_def)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   890
  then show ?thesis
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   891
    by (auto simp: space_restrict_space measure_restrict_space[OF \<Omega>(1)] le_infI2
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   892
                   simple_bochner_integral_def Collect_restrict
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   893
             split: split_indicator split_indicator_asm
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   894
             intro!: sum.mono_neutral_cong_left arg_cong2[where f=measure])
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   895
qed
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   896
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61609
diff changeset
   897
context
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 62083
diff changeset
   898
  notes [[inductive_internals]]
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61609
diff changeset
   899
begin
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61609
diff changeset
   900
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   901
inductive integrable for M f where
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   902
  "has_bochner_integral M f x \<Longrightarrow> integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   903
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61609
diff changeset
   904
end
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61609
diff changeset
   905
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69739
diff changeset
   906
definition\<^marker>\<open>tag important\<close> lebesgue_integral ("integral\<^sup>L") where
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
   907
  "integral\<^sup>L M f = (if \<exists>x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   908
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   909
syntax
59357
f366643536cd allow line breaks in integral notation
Andreas Lochbihler
parents: 59353
diff changeset
   910
  "_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> 'a measure \<Rightarrow> real" ("\<integral>((2 _./ _)/ \<partial>_)" [60,61] 110)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   911
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   912
translations
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   913
  "\<integral> x. f \<partial>M" == "CONST lebesgue_integral M (\<lambda>x. f)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   914
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   915
syntax
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   916
  "_ascii_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real" ("(3LINT (1_)/|(_)./ _)" [0,110,60] 60)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   917
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   918
translations
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   919
  "LINT x|M. f" == "CONST lebesgue_integral M (\<lambda>x. f)"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   920
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   921
lemma has_bochner_integral_integral_eq: "has_bochner_integral M f x \<Longrightarrow> integral\<^sup>L M f = x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   922
  by (metis the_equality has_bochner_integral_eq lebesgue_integral_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   923
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   924
lemma has_bochner_integral_integrable:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   925
  "integrable M f \<Longrightarrow> has_bochner_integral M f (integral\<^sup>L M f)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   926
  by (auto simp: has_bochner_integral_integral_eq integrable.simps)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   927
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   928
lemma has_bochner_integral_iff:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   929
  "has_bochner_integral M f x \<longleftrightarrow> integrable M f \<and> integral\<^sup>L M f = x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   930
  by (metis has_bochner_integral_integrable has_bochner_integral_integral_eq integrable.intros)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   931
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   932
lemma simple_bochner_integrable_eq_integral:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   933
  "simple_bochner_integrable M f \<Longrightarrow> simple_bochner_integral M f = integral\<^sup>L M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   934
  using has_bochner_integral_simple_bochner_integrable[of M f]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   935
  by (simp add: has_bochner_integral_integral_eq)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   936
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   937
lemma not_integrable_integral_eq: "\<not> integrable M f \<Longrightarrow> integral\<^sup>L M f = 0"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   938
  unfolding integrable.simps lebesgue_integral_def by (auto intro!: arg_cong[where f=The])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   939
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   940
lemma integral_eq_cases:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   941
  "integrable M f \<longleftrightarrow> integrable N g \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   942
    (integrable M f \<Longrightarrow> integrable N g \<Longrightarrow> integral\<^sup>L M f = integral\<^sup>L N g) \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   943
    integral\<^sup>L M f = integral\<^sup>L N g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   944
  by (metis not_integrable_integral_eq)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   945
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   946
lemma borel_measurable_integrable[measurable_dest]: "integrable M f \<Longrightarrow> f \<in> borel_measurable M"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   947
  by (auto elim: integrable.cases has_bochner_integral.cases)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   948
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   949
lemma borel_measurable_integrable'[measurable_dest]:
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   950
  "integrable M f \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N"
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   951
  using borel_measurable_integrable[measurable] by measurable
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   952
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   953
lemma integrable_cong:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   954
  "M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> integrable N g"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
   955
  by (simp cong: has_bochner_integral_cong add: integrable.simps)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   956
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   957
lemma integrable_cong_AE:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   958
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. f x = g x \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   959
    integrable M f \<longleftrightarrow> integrable M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   960
  unfolding integrable.simps
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   961
  by (intro has_bochner_integral_cong_AE arg_cong[where f=Ex] ext)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   962
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   963
lemma integrable_cong_AE_imp:
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
   964
  "integrable M g \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> (AE x in M. g x = f x) \<Longrightarrow> integrable M f"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
   965
  using integrable_cong_AE[of f M g] by (auto simp: eq_commute)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
   966
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   967
lemma integral_cong:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   968
  "M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integral\<^sup>L M f = integral\<^sup>L N g"
63566
e5abbdee461a more accurate cong del;
wenzelm
parents: 63092
diff changeset
   969
  by (simp cong: has_bochner_integral_cong cong del: if_weak_cong add: lebesgue_integral_def)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   970
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   971
lemma integral_cong_AE:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   972
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. f x = g x \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   973
    integral\<^sup>L M f = integral\<^sup>L M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   974
  unfolding lebesgue_integral_def
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
   975
  by (rule arg_cong[where x="has_bochner_integral M f"]) (intro has_bochner_integral_cong_AE ext)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   976
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   977
lemma integrable_add[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x + g x)"
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
   978
  by (auto simp: integrable.simps)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   979
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   980
lemma integrable_zero[simp, intro]: "integrable M (\<lambda>x. 0)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   981
  by (metis has_bochner_integral_zero integrable.simps)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   982
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   983
lemma integrable_sum[simp, intro]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> integrable M (\<lambda>x. \<Sum>i\<in>I. f i x)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   984
  by (metis has_bochner_integral_sum integrable.simps)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   985
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   986
lemma integrable_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   987
  integrable M (\<lambda>x. indicator A x *\<^sub>R c)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   988
  by (metis has_bochner_integral_indicator integrable.simps)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   989
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   990
lemma integrable_real_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   991
  integrable M (indicator A :: 'a \<Rightarrow> real)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   992
  by (metis has_bochner_integral_real_indicator integrable.simps)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   993
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   994
lemma integrable_diff[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x - g x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   995
  by (auto simp: integrable.simps intro: has_bochner_integral_diff)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   996
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
   997
lemma integrable_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. T (f x))"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   998
  by (auto simp: integrable.simps intro: has_bochner_integral_bounded_linear)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
   999
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1000
lemma integrable_scaleR_left[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x *\<^sub>R c)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1001
  unfolding integrable.simps by fastforce
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1002
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1003
lemma integrable_scaleR_right[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c *\<^sub>R f x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1004
  unfolding integrable.simps by fastforce
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1005
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1006
lemma integrable_mult_left[simp, intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1007
  fixes c :: "_::{real_normed_algebra,second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1008
  shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x * c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1009
  unfolding integrable.simps by fastforce
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1010
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1011
lemma integrable_mult_right[simp, intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1012
  fixes c :: "_::{real_normed_algebra,second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1013
  shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c * f x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1014
  unfolding integrable.simps by fastforce
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1015
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1016
lemma integrable_divide_zero[simp, intro]:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59587
diff changeset
  1017
  fixes c :: "_::{real_normed_field, field, second_countable_topology}"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1018
  shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x / c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1019
  unfolding integrable.simps by fastforce
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1020
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1021
lemma integrable_inner_left[simp, intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1022
  "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x \<bullet> c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1023
  unfolding integrable.simps by fastforce
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1024
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1025
lemma integrable_inner_right[simp, intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1026
  "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c \<bullet> f x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1027
  unfolding integrable.simps by fastforce
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1028
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1029
lemmas integrable_minus[simp, intro] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1030
  integrable_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1031
lemmas integrable_divide[simp, intro] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1032
  integrable_bounded_linear[OF bounded_linear_divide]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1033
lemmas integrable_Re[simp, intro] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1034
  integrable_bounded_linear[OF bounded_linear_Re]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1035
lemmas integrable_Im[simp, intro] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1036
  integrable_bounded_linear[OF bounded_linear_Im]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1037
lemmas integrable_cnj[simp, intro] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1038
  integrable_bounded_linear[OF bounded_linear_cnj]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1039
lemmas integrable_of_real[simp, intro] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1040
  integrable_bounded_linear[OF bounded_linear_of_real]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1041
lemmas integrable_fst[simp, intro] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1042
  integrable_bounded_linear[OF bounded_linear_fst]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1043
lemmas integrable_snd[simp, intro] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1044
  integrable_bounded_linear[OF bounded_linear_snd]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1045
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1046
lemma integral_zero[simp]: "integral\<^sup>L M (\<lambda>x. 0) = 0"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1047
  by (intro has_bochner_integral_integral_eq has_bochner_integral_zero)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1048
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1049
lemma integral_add[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1050
    integral\<^sup>L M (\<lambda>x. f x + g x) = integral\<^sup>L M f + integral\<^sup>L M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1051
  by (intro has_bochner_integral_integral_eq has_bochner_integral_add has_bochner_integral_integrable)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1052
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1053
lemma integral_diff[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1054
    integral\<^sup>L M (\<lambda>x. f x - g x) = integral\<^sup>L M f - integral\<^sup>L M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1055
  by (intro has_bochner_integral_integral_eq has_bochner_integral_diff has_bochner_integral_integrable)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1056
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1057
lemma integral_sum: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1058
  integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1059
  by (intro has_bochner_integral_integral_eq has_bochner_integral_sum has_bochner_integral_integrable)
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1060
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1061
lemma integral_sum'[simp]: "(\<And>i. i \<in> I =simp=> integrable M (f i)) \<Longrightarrow>
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
  1062
  integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1063
  unfolding simp_implies_def by (rule integral_sum)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
  1064
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1065
lemma integral_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1066
    integral\<^sup>L M (\<lambda>x. T (f x)) = T (integral\<^sup>L M f)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1067
  by (metis has_bochner_integral_bounded_linear has_bochner_integral_integrable has_bochner_integral_integral_eq)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1068
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1069
lemma integral_bounded_linear':
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1070
  assumes T: "bounded_linear T" and T': "bounded_linear T'"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1071
  assumes *: "\<not> (\<forall>x. T x = 0) \<Longrightarrow> (\<forall>x. T' (T x) = x)"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1072
  shows "integral\<^sup>L M (\<lambda>x. T (f x)) = T (integral\<^sup>L M f)"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1073
proof cases
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1074
  assume "(\<forall>x. T x = 0)" then show ?thesis
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1075
    by simp
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1076
next
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1077
  assume **: "\<not> (\<forall>x. T x = 0)"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1078
  show ?thesis
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1079
  proof cases
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1080
    assume "integrable M f" with T show ?thesis
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1081
      by (rule integral_bounded_linear)
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1082
  next
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1083
    assume not: "\<not> integrable M f"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1084
    moreover have "\<not> integrable M (\<lambda>x. T (f x))"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1085
    proof
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1086
      assume "integrable M (\<lambda>x. T (f x))"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1087
      from integrable_bounded_linear[OF T' this] not *[OF **]
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1088
      show False
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1089
        by auto
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1090
    qed
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1091
    ultimately show ?thesis
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1092
      using T by (simp add: not_integrable_integral_eq linear_simps)
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1093
  qed
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1094
qed
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1095
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1096
lemma integral_scaleR_left[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x *\<^sub>R c \<partial>M) = integral\<^sup>L M f *\<^sub>R c"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1097
  by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_scaleR_left)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1098
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1099
lemma integral_scaleR_right[simp]: "(\<integral> x. c *\<^sub>R f x \<partial>M) = c *\<^sub>R integral\<^sup>L M f"
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1100
  by (rule integral_bounded_linear'[OF bounded_linear_scaleR_right bounded_linear_scaleR_right[of "1 / c"]]) simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1101
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1102
lemma integral_mult_left[simp]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1103
  fixes c :: "_::{real_normed_algebra,second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1104
  shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x * c \<partial>M) = integral\<^sup>L M f * c"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1105
  by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_left)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1106
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1107
lemma integral_mult_right[simp]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1108
  fixes c :: "_::{real_normed_algebra,second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1109
  shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. c * f x \<partial>M) = c * integral\<^sup>L M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1110
  by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_right)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1111
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1112
lemma integral_mult_left_zero[simp]:
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1113
  fixes c :: "_::{real_normed_field,second_countable_topology}"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1114
  shows "(\<integral> x. f x * c \<partial>M) = integral\<^sup>L M f * c"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1115
  by (rule integral_bounded_linear'[OF bounded_linear_mult_left bounded_linear_mult_left[of "1 / c"]]) simp
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1116
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1117
lemma integral_mult_right_zero[simp]:
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1118
  fixes c :: "_::{real_normed_field,second_countable_topology}"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1119
  shows "(\<integral> x. c * f x \<partial>M) = c * integral\<^sup>L M f"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1120
  by (rule integral_bounded_linear'[OF bounded_linear_mult_right bounded_linear_mult_right[of "1 / c"]]) simp
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1121
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1122
lemma integral_inner_left[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x \<bullet> c \<partial>M) = integral\<^sup>L M f \<bullet> c"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1123
  by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_left)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1124
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1125
lemma integral_inner_right[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. c \<bullet> f x \<partial>M) = c \<bullet> integral\<^sup>L M f"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1126
  by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_right)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1127
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1128
lemma integral_divide_zero[simp]:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59587
diff changeset
  1129
  fixes c :: "_::{real_normed_field, field, second_countable_topology}"
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1130
  shows "integral\<^sup>L M (\<lambda>x. f x / c) = integral\<^sup>L M f / c"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1131
  by (rule integral_bounded_linear'[OF bounded_linear_divide bounded_linear_mult_left[of c]]) simp
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1132
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1133
lemma integral_minus[simp]: "integral\<^sup>L M (\<lambda>x. - f x) = - integral\<^sup>L M f"
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1134
  by (rule integral_bounded_linear'[OF bounded_linear_minus[OF bounded_linear_ident] bounded_linear_minus[OF bounded_linear_ident]]) simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1135
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1136
lemma integral_complex_of_real[simp]: "integral\<^sup>L M (\<lambda>x. complex_of_real (f x)) = of_real (integral\<^sup>L M f)"
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1137
  by (rule integral_bounded_linear'[OF bounded_linear_of_real bounded_linear_Re]) simp
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1138
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1139
lemma integral_cnj[simp]: "integral\<^sup>L M (\<lambda>x. cnj (f x)) = cnj (integral\<^sup>L M f)"
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1140
  by (rule integral_bounded_linear'[OF bounded_linear_cnj bounded_linear_cnj]) simp
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1141
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1142
lemmas integral_divide[simp] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1143
  integral_bounded_linear[OF bounded_linear_divide]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1144
lemmas integral_Re[simp] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1145
  integral_bounded_linear[OF bounded_linear_Re]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1146
lemmas integral_Im[simp] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1147
  integral_bounded_linear[OF bounded_linear_Im]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1148
lemmas integral_of_real[simp] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1149
  integral_bounded_linear[OF bounded_linear_of_real]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1150
lemmas integral_fst[simp] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1151
  integral_bounded_linear[OF bounded_linear_fst]
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1152
lemmas integral_snd[simp] =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1153
  integral_bounded_linear[OF bounded_linear_snd]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1154
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1155
lemma integral_norm_bound_ennreal:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1156
  "integrable M f \<Longrightarrow> norm (integral\<^sup>L M f) \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1157
  by (metis has_bochner_integral_integrable has_bochner_integral_norm_bound)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1158
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1159
lemma integrableI_sequence:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1160
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1161
  assumes f[measurable]: "f \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1162
  assumes s: "\<And>i. simple_bochner_integrable M (s i)"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1163
  assumes lim: "(\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?S \<longlonglongrightarrow> 0")
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1164
  shows "integrable M f"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1165
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1166
  let ?s = "\<lambda>n. simple_bochner_integral M (s n)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1167
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1168
  have "\<exists>x. ?s \<longlonglongrightarrow> x"
64287
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64283
diff changeset
  1169
    unfolding convergent_eq_Cauchy
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1170
  proof (rule metric_CauchyI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1171
    fix e :: real assume "0 < e"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1172
    then have "0 < ennreal (e / 2)" by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1173
    from order_tendstoD(2)[OF lim this]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1174
    obtain M where M: "\<And>n. M \<le> n \<Longrightarrow> ?S n < e / 2"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1175
      by (auto simp: eventually_sequentially)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1176
    show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (?s m) (?s n) < e"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1177
    proof (intro exI allI impI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1178
      fix m n assume m: "M \<le> m" and n: "M \<le> n"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1179
      have "?S n \<noteq> \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1180
        using M[OF n] by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1181
      have "norm (?s n - ?s m) \<le> ?S n + ?S m"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1182
        by (intro simple_bochner_integral_bounded s f)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1183
      also have "\<dots> < ennreal (e / 2) + e / 2"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1184
        by (intro add_strict_mono M n m)
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68073
diff changeset
  1185
      also have "\<dots> = e" using \<open>0<e\<close> by (simp flip: ennreal_plus)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1186
      finally show "dist (?s n) (?s m) < e"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1187
        using \<open>0<e\<close> by (simp add: dist_norm ennreal_less_iff)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1188
    qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1189
  qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1190
  then obtain x where "?s \<longlonglongrightarrow> x" ..
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1191
  show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1192
    by (rule, rule) fact+
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1193
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1194
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1195
proposition nn_integral_dominated_convergence_norm:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1196
  fixes u' :: "_ \<Rightarrow> _::{real_normed_vector, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1197
  assumes [measurable]:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1198
       "\<And>i. u i \<in> borel_measurable M" "u' \<in> borel_measurable M" "w \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1199
    and bound: "\<And>j. AE x in M. norm (u j x) \<le> w x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1200
    and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1201
    and u': "AE x in M. (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1202
  shows "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) \<longlonglongrightarrow> 0"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1203
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1204
  have "AE x in M. \<forall>j. norm (u j x) \<le> w x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1205
    unfolding AE_all_countable by rule fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1206
  with u' have bnd: "AE x in M. \<forall>j. norm (u' x - u j x) \<le> 2 * w x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1207
  proof (eventually_elim, intro allI)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1208
    fix i x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x" "\<forall>j. norm (u j x) \<le> w x" "\<forall>j. norm (u j x) \<le> w x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1209
    then have "norm (u' x) \<le> w x" "norm (u i x) \<le> w x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1210
      by (auto intro: LIMSEQ_le_const2 tendsto_norm)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1211
    then have "norm (u' x) + norm (u i x) \<le> 2 * w x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1212
      by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1213
    also have "norm (u' x - u i x) \<le> norm (u' x) + norm (u i x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1214
      by (rule norm_triangle_ineq4)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1215
    finally (xtrans) show "norm (u' x - u i x) \<le> 2 * w x" .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1216
  qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1217
  have w_nonneg: "AE x in M. 0 \<le> w x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1218
    using bound[of 0] by (auto intro: order_trans[OF norm_ge_zero])
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1219
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1220
  have "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. 0 \<partial>M)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1221
  proof (rule nn_integral_dominated_convergence)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1222
    show "(\<integral>\<^sup>+x. 2 * w x \<partial>M) < \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1223
      by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) (insert w_nonneg, auto simp: ennreal_mult )
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1224
    show "AE x in M. (\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1225
      using u'
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1226
    proof eventually_elim
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1227
      fix x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1228
      from tendsto_diff[OF tendsto_const[of "u' x"] this]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1229
      show "(\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68073
diff changeset
  1230
        by (simp add: tendsto_norm_zero_iff flip: ennreal_0)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1231
    qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1232
  qed (insert bnd w_nonneg, auto)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1233
  then show ?thesis by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1234
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1235
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1236
proposition integrableI_bounded:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1237
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1238
  assumes f[measurable]: "f \<in> borel_measurable M" and fin: "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1239
  shows "integrable M f"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1240
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1241
  from borel_measurable_implies_sequence_metric[OF f, of 0] obtain s where
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1242
    s: "\<And>i. simple_function M (s i)" and
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1243
    pointwise: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" and
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1244
    bound: "\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)"
62379
340738057c8c An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents: 62093
diff changeset
  1245
    by simp metis
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1246
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1247
  show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1248
  proof (rule integrableI_sequence)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1249
    { fix i
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1250
      have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1251
        by (intro nn_integral_mono) (simp add: bound)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1252
      also have "\<dots> = 2 * (\<integral>\<^sup>+x. ennreal (norm (f x)) \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1253
        by (simp add: ennreal_mult nn_integral_cmult)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1254
      also have "\<dots> < top"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1255
        using fin by (simp add: ennreal_mult_less_top)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1256
      finally have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1257
        by simp }
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1258
    note fin_s = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1259
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1260
    show "\<And>i. simple_bochner_integrable M (s i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1261
      by (rule simple_bochner_integrableI_bounded) fact+
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1262
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1263
    show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1264
    proof (rule nn_integral_dominated_convergence_norm)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1265
      show "\<And>j. AE x in M. norm (s j x) \<le> 2 * norm (f x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1266
        using bound by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1267
      show "\<And>i. s i \<in> borel_measurable M" "(\<lambda>x. 2 * norm (f x)) \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1268
        using s by (auto intro: borel_measurable_simple_function)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1269
      show "(\<integral>\<^sup>+ x. ennreal (2 * norm (f x)) \<partial>M) < \<infinity>"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1270
        using fin by (simp add: nn_integral_cmult ennreal_mult ennreal_mult_less_top)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1271
      show "AE x in M. (\<lambda>i. s i x) \<longlonglongrightarrow> f x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1272
        using pointwise by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1273
    qed fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1274
  qed fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1275
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1276
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1277
lemma integrableI_bounded_set:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1278
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1279
  assumes [measurable]: "A \<in> sets M" "f \<in> borel_measurable M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1280
  assumes finite: "emeasure M A < \<infinity>"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1281
    and bnd: "AE x in M. x \<in> A \<longrightarrow> norm (f x) \<le> B"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1282
    and null: "AE x in M. x \<notin> A \<longrightarrow> f x = 0"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1283
  shows "integrable M f"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1284
proof (rule integrableI_bounded)
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1285
  { fix x :: 'b have "norm x \<le> B \<Longrightarrow> 0 \<le> B"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1286
      using norm_ge_zero[of x] by arith }
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1287
  with bnd null have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (max 0 B) * indicator A x \<partial>M)"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1288
    by (intro nn_integral_mono_AE) (auto split: split_indicator split_max)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1289
  also have "\<dots> < \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1290
    using finite by (subst nn_integral_cmult_indicator) (auto simp: ennreal_mult_less_top)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1291
  finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1292
qed simp
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1293
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1294
lemma integrableI_bounded_set_indicator:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1295
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1296
  shows "A \<in> sets M \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow>
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1297
    emeasure M A < \<infinity> \<Longrightarrow> (AE x in M. x \<in> A \<longrightarrow> norm (f x) \<le> B) \<Longrightarrow>
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1298
    integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1299
  by (rule integrableI_bounded_set[where A=A]) auto
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1300
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1301
lemma integrableI_nonneg:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1302
  fixes f :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1303
  assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1304
  shows "integrable M f"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1305
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1306
  have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1307
    using assms by (intro nn_integral_cong_AE) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1308
  then show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1309
    using assms by (intro integrableI_bounded) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1310
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1311
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1312
lemma integrable_iff_bounded:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1313
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1314
  shows "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1315
  using integrableI_bounded[of f M] has_bochner_integral_implies_finite_norm[of M f]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1316
  unfolding integrable.simps has_bochner_integral.simps[abs_def] by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1317
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1318
lemma (in finite_measure) square_integrable_imp_integrable:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1319
  fixes f :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, real_normed_div_algebra}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1320
  assumes [measurable]: "f \<in> borel_measurable M"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1321
  assumes "integrable M (\<lambda>x. f x ^ 2)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1322
  shows   "integrable M f"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1323
proof -
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1324
  have less_top: "emeasure M (space M) < top"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1325
    using finite_emeasure_space by (meson top.not_eq_extremum)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1326
  have "nn_integral M (\<lambda>x. norm (f x)) ^ 2 \<le>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1327
        nn_integral M (\<lambda>x. norm (f x) ^ 2) * emeasure M (space M)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1328
    using Cauchy_Schwarz_nn_integral[of "\<lambda>x. norm (f x)" M "\<lambda>_. 1"]
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1329
    by (simp add: ennreal_power)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1330
  also have "\<dots> < \<infinity>"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1331
    using assms(2) less_top
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1332
    by (subst (asm) integrable_iff_bounded) (auto simp: norm_power ennreal_mult_less_top)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1333
  finally have "nn_integral M (\<lambda>x. norm (f x)) < \<infinity>"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1334
    by (simp add: power_less_top_ennreal)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1335
  thus ?thesis
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1336
    by (simp add: integrable_iff_bounded)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1337
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  1338
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1339
lemma integrable_bound:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1340
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1341
    and g :: "'a \<Rightarrow> 'c::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1342
  shows "integrable M f \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. norm (g x) \<le> norm (f x)) \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1343
    integrable M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1344
  unfolding integrable_iff_bounded
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1345
proof safe
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1346
  assume "f \<in> borel_measurable M" "g \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1347
  assume "AE x in M. norm (g x) \<le> norm (f x)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1348
  then have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1349
    by  (intro nn_integral_mono_AE) auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1350
  also assume "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1351
  finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>" .
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1352
qed
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1353
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1354
lemma integrable_mult_indicator:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1355
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1356
  shows "A \<in> sets M \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1357
  by (rule integrable_bound[of M f]) (auto split: split_indicator)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1358
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1359
lemma integrable_real_mult_indicator:
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1360
  fixes f :: "'a \<Rightarrow> real"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1361
  shows "A \<in> sets M \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. f x * indicator A x)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1362
  using integrable_mult_indicator[of A M f] by (simp add: mult_ac)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1363
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1364
lemma integrable_abs[simp, intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1365
  fixes f :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1366
  assumes [measurable]: "integrable M f" shows "integrable M (\<lambda>x. \<bar>f x\<bar>)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1367
  using assms by (rule integrable_bound) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1368
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1369
lemma integrable_norm[simp, intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1370
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1371
  assumes [measurable]: "integrable M f" shows "integrable M (\<lambda>x. norm (f x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1372
  using assms by (rule integrable_bound) auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1373
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1374
lemma integrable_norm_cancel:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1375
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1376
  assumes [measurable]: "integrable M (\<lambda>x. norm (f x))" "f \<in> borel_measurable M" shows "integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1377
  using assms by (rule integrable_bound) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1378
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1379
lemma integrable_norm_iff:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1380
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1381
  shows "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda>x. norm (f x)) \<longleftrightarrow> integrable M f"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1382
  by (auto intro: integrable_norm_cancel)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1383
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1384
lemma integrable_abs_cancel:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1385
  fixes f :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1386
  assumes [measurable]: "integrable M (\<lambda>x. \<bar>f x\<bar>)" "f \<in> borel_measurable M" shows "integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1387
  using assms by (rule integrable_bound) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1388
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1389
lemma integrable_abs_iff:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1390
  fixes f :: "'a \<Rightarrow> real"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1391
  shows "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda>x. \<bar>f x\<bar>) \<longleftrightarrow> integrable M f"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1392
  by (auto intro: integrable_abs_cancel)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1393
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1394
lemma integrable_max[simp, intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1395
  fixes f :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1396
  assumes fg[measurable]: "integrable M f" "integrable M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1397
  shows "integrable M (\<lambda>x. max (f x) (g x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1398
  using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1399
  by (rule integrable_bound) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1400
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1401
lemma integrable_min[simp, intro]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1402
  fixes f :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1403
  assumes fg[measurable]: "integrable M f" "integrable M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1404
  shows "integrable M (\<lambda>x. min (f x) (g x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1405
  using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1406
  by (rule integrable_bound) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1407
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1408
lemma integral_minus_iff[simp]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1409
  "integrable M (\<lambda>x. - f x ::'a::{banach, second_countable_topology}) \<longleftrightarrow> integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1410
  unfolding integrable_iff_bounded
71633
07bec530f02e cleaned proofs
nipkow
parents: 70532
diff changeset
  1411
  by (auto)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1412
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1413
lemma integrable_indicator_iff:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1414
  "integrable M (indicator A::_ \<Rightarrow> real) \<longleftrightarrow> A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1415
  by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator nn_integral_indicator'
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1416
           cong: conj_cong)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1417
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1418
lemma integral_indicator[simp]: "integral\<^sup>L M (indicator A) = measure M (A \<inter> space M)"
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1419
proof cases
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1420
  assume *: "A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1421
  have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \<inter> space M))"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1422
    by (intro integral_cong) (auto split: split_indicator)
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1423
  also have "\<dots> = measure M (A \<inter> space M)"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1424
    using * by (intro has_bochner_integral_integral_eq has_bochner_integral_real_indicator) auto
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1425
  finally show ?thesis .
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1426
next
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1427
  assume *: "\<not> (A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>)"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1428
  have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \<inter> space M) :: _ \<Rightarrow> real)"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1429
    by (intro integral_cong) (auto split: split_indicator)
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1430
  also have "\<dots> = 0"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1431
    using * by (subst not_integrable_integral_eq) (auto simp: integrable_indicator_iff)
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1432
  also have "\<dots> = measure M (A \<inter> space M)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1433
    using * by (auto simp: measure_def emeasure_notin_sets not_less top_unique)
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1434
  finally show ?thesis .
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1435
qed
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1436
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1437
lemma integrable_discrete_difference:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1438
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1439
  assumes X: "countable X"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1440
  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1441
  assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1442
  assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1443
  shows "integrable M f \<longleftrightarrow> integrable M g"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1444
  unfolding integrable_iff_bounded
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1445
proof (rule conj_cong)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1446
  { assume "f \<in> borel_measurable M" then have "g \<in> borel_measurable M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1447
      by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) }
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1448
  moreover
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1449
  { assume "g \<in> borel_measurable M" then have "f \<in> borel_measurable M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1450
      by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) }
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1451
  ultimately show "f \<in> borel_measurable M \<longleftrightarrow> g \<in> borel_measurable M" ..
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1452
next
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1453
  have "AE x in M. x \<notin> X"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1454
    by (rule AE_discrete_difference) fact+
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1455
  then have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. norm (g x) \<partial>M)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1456
    by (intro nn_integral_cong_AE) (auto simp: eq)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1457
  then show "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) < \<infinity> \<longleftrightarrow> (\<integral>\<^sup>+ x. norm (g x) \<partial>M) < \<infinity>"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1458
    by simp
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1459
qed
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1460
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1461
lemma integral_discrete_difference:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1462
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1463
  assumes X: "countable X"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1464
  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1465
  assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1466
  assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1467
  shows "integral\<^sup>L M f = integral\<^sup>L M g"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1468
proof (rule integral_eq_cases)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1469
  show eq: "integrable M f \<longleftrightarrow> integrable M g"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1470
    by (rule integrable_discrete_difference[where X=X]) fact+
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1471
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1472
  assume f: "integrable M f"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1473
  show "integral\<^sup>L M f = integral\<^sup>L M g"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1474
  proof (rule integral_cong_AE)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1475
    show "f \<in> borel_measurable M" "g \<in> borel_measurable M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1476
      using f eq by (auto intro: borel_measurable_integrable)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1477
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1478
    have "AE x in M. x \<notin> X"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1479
      by (rule AE_discrete_difference) fact+
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1480
    with AE_space show "AE x in M. f x = g x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1481
      by eventually_elim fact
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1482
  qed
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1483
qed
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1484
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1485
lemma has_bochner_integral_discrete_difference:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1486
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1487
  assumes X: "countable X"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1488
  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1489
  assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1490
  assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1491
  shows "has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral M g x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1492
  using integrable_discrete_difference[of X M f g, OF assms]
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1493
  using integral_discrete_difference[of X M f g, OF assms]
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1494
  by (metis has_bochner_integral_iff)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1495
69739
nipkow
parents: 69722
diff changeset
  1496
lemma
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1497
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and w :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1498
  assumes "f \<in> borel_measurable M" "\<And>i. s i \<in> borel_measurable M" "integrable M w"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1499
  assumes lim: "AE x in M. (\<lambda>i. s i x) \<longlonglongrightarrow> f x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1500
  assumes bound: "\<And>i. AE x in M. norm (s i x) \<le> w x"
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  1501
  shows integrable_dominated_convergence: "integrable M f"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  1502
    and integrable_dominated_convergence2: "\<And>i. integrable M (s i)"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1503
    and integral_dominated_convergence: "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1504
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1505
  have w_nonneg: "AE x in M. 0 \<le> w x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1506
    using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1507
  then have "(\<integral>\<^sup>+x. w x \<partial>M) = (\<integral>\<^sup>+x. norm (w x) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1508
    by (intro nn_integral_cong_AE) auto
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
  1509
  with \<open>integrable M w\<close> have w: "w \<in> borel_measurable M" "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1510
    unfolding integrable_iff_bounded by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1511
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1512
  show int_s: "\<And>i. integrable M (s i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1513
    unfolding integrable_iff_bounded
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1514
  proof
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1515
    fix i
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1516
    have "(\<integral>\<^sup>+ x. ennreal (norm (s i x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1517
      using bound[of i] w_nonneg by (intro nn_integral_mono_AE) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1518
    with w show "(\<integral>\<^sup>+ x. ennreal (norm (s i x)) \<partial>M) < \<infinity>" by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1519
  qed fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1520
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1521
  have all_bound: "AE x in M. \<forall>i. norm (s i x) \<le> w x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1522
    using bound unfolding AE_all_countable by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1523
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1524
  show int_f: "integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1525
    unfolding integrable_iff_bounded
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1526
  proof
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1527
    have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1528
      using all_bound lim w_nonneg
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1529
    proof (intro nn_integral_mono_AE, eventually_elim)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1530
      fix x assume "\<forall>i. norm (s i x) \<le> w x" "(\<lambda>i. s i x) \<longlonglongrightarrow> f x" "0 \<le> w x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1531
      then show "ennreal (norm (f x)) \<le> ennreal (w x)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1532
        by (intro LIMSEQ_le_const2[where X="\<lambda>i. ennreal (norm (s i x))"]) (auto intro: tendsto_intros)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1533
    qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1534
    with w show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1535
  qed fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1536
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1537
  have "(\<lambda>n. ennreal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) \<longlonglongrightarrow> ennreal 0" (is "?d \<longlonglongrightarrow> ennreal 0")
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1538
  proof (rule tendsto_sandwich)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1539
    show "eventually (\<lambda>n. ennreal 0 \<le> ?d n) sequentially" "(\<lambda>_. ennreal 0) \<longlonglongrightarrow> ennreal 0" by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1540
    show "eventually (\<lambda>n. ?d n \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)) sequentially"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1541
    proof (intro always_eventually allI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1542
      fix n
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1543
      have "?d n = norm (integral\<^sup>L M (\<lambda>x. s n x - f x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1544
        using int_f int_s by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1545
      also have "\<dots> \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1546
        by (intro int_f int_s integrable_diff integral_norm_bound_ennreal)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1547
      finally show "?d n \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)" .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1548
    qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1549
    show "(\<lambda>n. \<integral>\<^sup>+x. norm (s n x - f x) \<partial>M) \<longlonglongrightarrow> ennreal 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1550
      unfolding ennreal_0
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1551
      apply (subst norm_minus_commute)
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1552
    proof (rule nn_integral_dominated_convergence_norm[where w=w])
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1553
      show "\<And>n. s n \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1554
        using int_s unfolding integrable_iff_bounded by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1555
    qed fact+
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1556
  qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1557
  then have "(\<lambda>n. integral\<^sup>L M (s n) - integral\<^sup>L M f) \<longlonglongrightarrow> 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1558
    by (simp add: tendsto_norm_zero_iff del: ennreal_0)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1559
  from tendsto_add[OF this tendsto_const[of "integral\<^sup>L M f"]]
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1560
  show "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f"  by simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1561
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1562
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1563
context
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1564
  fixes s :: "real \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}" and w :: "'a \<Rightarrow> real"
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1565
    and f :: "'a \<Rightarrow> 'b" and M
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1566
  assumes "f \<in> borel_measurable M" "\<And>t. s t \<in> borel_measurable M" "integrable M w"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1567
  assumes lim: "AE x in M. ((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top"
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1568
  assumes bound: "\<forall>\<^sub>F i in at_top. AE x in M. norm (s i x) \<le> w x"
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1569
begin
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1570
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1571
lemma integral_dominated_convergence_at_top: "((\<lambda>t. integral\<^sup>L M (s t)) \<longlongrightarrow> integral\<^sup>L M f) at_top"
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1572
proof (rule tendsto_at_topI_sequentially)
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1573
  fix X :: "nat \<Rightarrow> real" assume X: "filterlim X at_top sequentially"
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1574
  from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound]
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1575
  obtain N where w: "\<And>n. N \<le> n \<Longrightarrow> AE x in M. norm (s (X n) x) \<le> w x"
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1576
    by (auto simp: eventually_sequentially)
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1577
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1578
  show "(\<lambda>n. integral\<^sup>L M (s (X n))) \<longlonglongrightarrow> integral\<^sup>L M f"
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1579
  proof (rule LIMSEQ_offset, rule integral_dominated_convergence)
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1580
    show "AE x in M. norm (s (X (n + N)) x) \<le> w x" for n
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1581
      by (rule w) auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1582
    show "AE x in M. (\<lambda>n. s (X (n + N)) x) \<longlonglongrightarrow> f x"
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1583
      using lim
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1584
    proof eventually_elim
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1585
      fix x assume "((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1586
      then show "(\<lambda>n. s (X (n + N)) x) \<longlonglongrightarrow> f x"
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1587
        by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X])
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1588
    qed
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1589
  qed fact+
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1590
qed
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1591
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1592
lemma integrable_dominated_convergence_at_top: "integrable M f"
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1593
proof -
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1594
  from bound obtain N where w: "\<And>n. N \<le> n \<Longrightarrow> AE x in M. norm (s n x) \<le> w x"
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1595
    by (auto simp: eventually_at_top_linorder)
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1596
  show ?thesis
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1597
  proof (rule integrable_dominated_convergence)
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1598
    show "AE x in M. norm (s (N + i) x) \<le> w x" for i :: nat
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1599
      by (intro w) auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1600
    show "AE x in M. (\<lambda>i. s (N + real i) x) \<longlonglongrightarrow> f x"
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1601
      using lim
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1602
    proof eventually_elim
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1603
      fix x assume "((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1604
      then show "(\<lambda>n. s (N + n) x) \<longlonglongrightarrow> f x"
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1605
        by (rule filterlim_compose)
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1606
           (auto intro!: filterlim_tendsto_add_at_top filterlim_real_sequentially)
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1607
    qed
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1608
  qed fact+
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1609
qed
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1610
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1611
end
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61880
diff changeset
  1612
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1613
lemma integrable_mult_left_iff [simp]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1614
  fixes f :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1615
  shows "integrable M (\<lambda>x. c * f x) \<longleftrightarrow> c = 0 \<or> integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1616
  using integrable_mult_left[of c M f] integrable_mult_left[of "1 / c" M "\<lambda>x. c * f x"]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1617
  by (cases "c = 0") auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1618
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1619
lemma integrable_mult_right_iff [simp]:
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1620
  fixes f :: "'a \<Rightarrow> real"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1621
  shows "integrable M (\<lambda>x. f x * c) \<longleftrightarrow> c = 0 \<or> integrable M f"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1622
  using integrable_mult_left_iff [of M c f] by (simp add: mult.commute)
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  1623
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1624
lemma integrableI_nn_integral_finite:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1625
  assumes [measurable]: "f \<in> borel_measurable M"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1626
    and nonneg: "AE x in M. 0 \<le> f x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1627
    and finite: "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1628
  shows "integrable M f"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1629
proof (rule integrableI_bounded)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1630
  have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1631
    using nonneg by (intro nn_integral_cong_AE) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1632
  with finite show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1633
    by auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1634
qed simp
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1635
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1636
lemma integral_nonneg_AE:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1637
  fixes f :: "'a \<Rightarrow> real"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1638
  assumes nonneg: "AE x in M. 0 \<le> f x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1639
  shows "0 \<le> integral\<^sup>L M f"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1640
proof cases
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1641
  assume f: "integrable M f"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1642
  then have [measurable]: "f \<in> M \<rightarrow>\<^sub>M borel"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1643
    by auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1644
  have "(\<lambda>x. max 0 (f x)) \<in> M \<rightarrow>\<^sub>M borel" "\<And>x. 0 \<le> max 0 (f x)" "integrable M (\<lambda>x. max 0 (f x))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1645
    using f by auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1646
  from this have "0 \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1647
  proof (induction rule: borel_measurable_induct_real)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1648
    case (add f g)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1649
    then have "integrable M f" "integrable M g"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1650
      by (auto intro!: integrable_bound[OF add.prems])
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1651
    with add show ?case
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1652
      by (simp add: nn_integral_add)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1653
  next
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1654
    case (seq U)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1655
    show ?case
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1656
    proof (rule LIMSEQ_le_const)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1657
      have U_le: "x \<in> space M \<Longrightarrow> U i x \<le> max 0 (f x)" for x i
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1658
        using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1659
      with seq nonneg show "(\<lambda>i. integral\<^sup>L M (U i)) \<longlonglongrightarrow> LINT x|M. max 0 (f x)"
73869
7181130f5872 more default simp rules
haftmann
parents: 73536
diff changeset
  1660
        by (intro integral_dominated_convergence)
7181130f5872 more default simp rules
haftmann
parents: 73536
diff changeset
  1661
          (simp_all, fastforce)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1662
      have "integrable M (U i)" for i
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1663
        using seq.prems by (rule integrable_bound) (insert U_le seq, auto)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1664
      with seq show "\<exists>N. \<forall>n\<ge>N. 0 \<le> integral\<^sup>L M (U n)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1665
        by auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1666
    qed
71633
07bec530f02e cleaned proofs
nipkow
parents: 70532
diff changeset
  1667
  qed (auto)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1668
  also have "\<dots> = integral\<^sup>L M f"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1669
    using nonneg by (auto intro!: integral_cong_AE)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1670
  finally show ?thesis .
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1671
qed (simp add: not_integrable_integral_eq)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1672
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1673
lemma integral_nonneg[simp]:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1674
  fixes f :: "'a \<Rightarrow> real"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1675
  shows "(\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> 0 \<le> integral\<^sup>L M f"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1676
  by (intro integral_nonneg_AE) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1677
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1678
proposition nn_integral_eq_integral:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1679
  assumes f: "integrable M f"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1680
  assumes nonneg: "AE x in M. 0 \<le> f x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1681
  shows "(\<integral>\<^sup>+ x. f x \<partial>M) = integral\<^sup>L M f"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1682
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1683
  { fix f :: "'a \<Rightarrow> real" assume f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" "integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1684
    then have "(\<integral>\<^sup>+ x. f x \<partial>M) = integral\<^sup>L M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1685
    proof (induct rule: borel_measurable_induct_real)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1686
      case (set A) then show ?case
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1687
        by (simp add: integrable_indicator_iff ennreal_indicator emeasure_eq_ennreal_measure)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1688
    next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1689
      case (mult f c) then show ?case
71633
07bec530f02e cleaned proofs
nipkow
parents: 70532
diff changeset
  1690
        by (auto simp add: nn_integral_cmult ennreal_mult integral_nonneg_AE)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1691
    next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1692
      case (add g f)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1693
      then have "integrable M f" "integrable M g"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1694
        by (auto intro!: integrable_bound[OF add.prems])
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1695
      with add show ?case
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1696
        by (simp add: nn_integral_add integral_nonneg_AE)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1697
    next
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1698
      case (seq U)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1699
      show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1700
      proof (rule LIMSEQ_unique)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1701
        have U_le_f: "x \<in> space M \<Longrightarrow> U i x \<le> f x" for x i
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1702
          using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1703
        have int_U: "\<And>i. integrable M (U i)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1704
          using seq f U_le_f by (intro integrable_bound[OF f(3)]) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1705
        from U_le_f seq have "(\<lambda>i. integral\<^sup>L M (U i)) \<longlonglongrightarrow> integral\<^sup>L M f"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1706
          by (intro integral_dominated_convergence) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1707
        then show "(\<lambda>i. ennreal (integral\<^sup>L M (U i))) \<longlonglongrightarrow> ennreal (integral\<^sup>L M f)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1708
          using seq f int_U by (simp add: f integral_nonneg_AE)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1709
        have "(\<lambda>i. \<integral>\<^sup>+ x. U i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+ x. f x \<partial>M"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1710
          using seq U_le_f f
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1711
          by (intro nn_integral_dominated_convergence[where w=f]) (auto simp: integrable_iff_bounded)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1712
        then show "(\<lambda>i. \<integral>x. U i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+x. f x \<partial>M"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1713
          using seq int_U by simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1714
      qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1715
    qed }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1716
  from this[of "\<lambda>x. max 0 (f x)"] assms have "(\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) = integral\<^sup>L M (\<lambda>x. max 0 (f x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1717
    by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1718
  also have "\<dots> = integral\<^sup>L M f"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1719
    using assms by (auto intro!: integral_cong_AE simp: integral_nonneg_AE)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1720
  also have "(\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1721
    using assms by (auto intro!: nn_integral_cong_AE simp: max_def)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1722
  finally show ?thesis .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1723
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1724
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1725
lemma nn_integral_eq_integrable:
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  1726
  assumes f: "f \<in> M \<rightarrow>\<^sub>M borel" "AE x in M. 0 \<le> f x" and "0 \<le> x"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  1727
  shows "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x \<longleftrightarrow> (integrable M f \<and> integral\<^sup>L M f = x)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  1728
proof (safe intro!: nn_integral_eq_integral assms)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  1729
  assume *: "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  1730
  with integrableI_nn_integral_finite[OF f this] nn_integral_eq_integral[of M f, OF _ f(2)]
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  1731
  show "integrable M f" "integral\<^sup>L M f = x"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  1732
    by (simp_all add: * assms integral_nonneg_AE)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  1733
qed
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  1734
69739
nipkow
parents: 69722
diff changeset
  1735
lemma
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1736
  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> 'a :: {banach, second_countable_topology}"
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1737
  assumes integrable[measurable]: "\<And>i. integrable M (f i)"
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1738
  and summable: "AE x in M. summable (\<lambda>i. norm (f i x))"
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1739
  and sums: "summable (\<lambda>i. (\<integral>x. norm (f i x) \<partial>M))"
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1740
  shows integrable_suminf: "integrable M (\<lambda>x. (\<Sum>i. f i x))" (is "integrable M ?S")
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1741
    and sums_integral: "(\<lambda>i. integral\<^sup>L M (f i)) sums (\<integral>x. (\<Sum>i. f i x) \<partial>M)" (is "?f sums ?x")
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1742
    and integral_suminf: "(\<integral>x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>L M (f i))"
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1743
    and summable_integral: "summable (\<lambda>i. integral\<^sup>L M (f i))"
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1744
proof -
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1745
  have 1: "integrable M (\<lambda>x. \<Sum>i. norm (f i x))"
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1746
  proof (rule integrableI_bounded)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1747
    have "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) = (\<integral>\<^sup>+ x. (\<Sum>i. ennreal (norm (f i x))) \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1748
      apply (intro nn_integral_cong_AE)
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1749
      using summable
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1750
      apply eventually_elim
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1751
      apply (simp add: suminf_nonneg ennreal_suminf_neq_top)
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1752
      done
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1753
    also have "\<dots> = (\<Sum>i. \<integral>\<^sup>+ x. norm (f i x) \<partial>M)"
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1754
      by (intro nn_integral_suminf) auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1755
    also have "\<dots> = (\<Sum>i. ennreal (\<integral>x. norm (f i x) \<partial>M))"
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1756
      by (intro arg_cong[where f=suminf] ext nn_integral_eq_integral integrable_norm integrable) auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1757
    finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) < \<infinity>"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1758
      by (simp add: sums ennreal_suminf_neq_top less_top[symmetric] integral_nonneg_AE)
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1759
  qed simp
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1760
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1761
  have 2: "AE x in M. (\<lambda>n. \<Sum>i<n. f i x) \<longlonglongrightarrow> (\<Sum>i. f i x)"
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1762
    using summable by eventually_elim (auto intro: summable_LIMSEQ summable_norm_cancel)
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1763
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1764
  have 3: "\<And>j. AE x in M. norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))"
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1765
    using summable
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1766
  proof eventually_elim
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1767
    fix j x assume [simp]: "summable (\<lambda>i. norm (f i x))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1768
    have "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i<j. norm (f i x))" by (rule norm_sum)
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1769
    also have "\<dots> \<le> (\<Sum>i. norm (f i x))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1770
      using sum_le_suminf[of "\<lambda>i. norm (f i x)"] unfolding sums_iff by auto
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1771
    finally show "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))" by simp
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1772
  qed
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1773
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  1774
  note ibl = integrable_dominated_convergence[OF _ _ 1 2 3]
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  1775
  note int = integral_dominated_convergence[OF _ _ 1 2 3]
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1776
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1777
  show "integrable M ?S"
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  1778
    by (rule ibl) measurable
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1779
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1780
  show "?f sums ?x" unfolding sums_def
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  1781
    using int by (simp add: integrable)
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1782
  then show "?x = suminf ?f" "summable ?f"
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1783
    unfolding sums_iff by auto
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1784
qed
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  1785
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1786
proposition integral_norm_bound [simp]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1787
  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1788
  shows "norm (integral\<^sup>L M f) \<le> (\<integral>x. norm (f x) \<partial>M)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1789
proof (cases "integrable M f")
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1790
  case True then show ?thesis
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1791
    using nn_integral_eq_integral[of M "\<lambda>x. norm (f x)"] integral_norm_bound_ennreal[of M f]
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1792
    by (simp add: integral_nonneg_AE)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1793
next
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1794
  case False
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1795
  then have "norm (integral\<^sup>L M f) = 0" by (simp add: not_integrable_integral_eq)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1796
  moreover have "(\<integral>x. norm (f x) \<partial>M) \<ge> 0" by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1797
  ultimately show ?thesis by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1798
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1799
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1800
proposition integral_abs_bound [simp]:
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1801
  fixes f :: "'a \<Rightarrow> real" shows "abs (\<integral>x. f x \<partial>M) \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1802
using integral_norm_bound[of M f] by auto
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1803
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1804
lemma integral_eq_nn_integral:
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1805
  assumes [measurable]: "f \<in> borel_measurable M"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1806
  assumes nonneg: "AE x in M. 0 \<le> f x"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1807
  shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)"
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1808
proof cases
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1809
  assume *: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = \<infinity>"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1810
  also have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1811
    using nonneg by (intro nn_integral_cong_AE) auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1812
  finally have "\<not> integrable M f"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1813
    by (auto simp: integrable_iff_bounded)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1814
  then show ?thesis
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1815
    by (simp add: * not_integrable_integral_eq)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1816
next
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1817
  assume "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>"
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1818
  then have "integrable M f"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1819
    by (cases "\<integral>\<^sup>+ x. ennreal (f x) \<partial>M" rule: ennreal_cases)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1820
       (auto intro!: integrableI_nn_integral_finite assms)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1821
  from nn_integral_eq_integral[OF this] nonneg show ?thesis
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1822
    by (simp add: integral_nonneg_AE)
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1823
qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1824
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1825
lemma enn2real_nn_integral_eq_integral:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1826
  assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \<le> g x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1827
    and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < top"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1828
    and [measurable]: "g \<in> M \<rightarrow>\<^sub>M borel"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1829
  shows "enn2real (\<integral>\<^sup>+x. f x \<partial>M) = (\<integral>x. g x \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1830
proof -
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1831
  have "ennreal (enn2real (\<integral>\<^sup>+x. f x \<partial>M)) = (\<integral>\<^sup>+x. f x \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1832
    using fin by (intro ennreal_enn2real) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1833
  also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1834
    using eq by (rule nn_integral_cong_AE)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1835
  also have "\<dots> = (\<integral>x. g x \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1836
  proof (rule nn_integral_eq_integral)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1837
    show "integrable M g"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1838
    proof (rule integrableI_bounded)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1839
      have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1840
        using eq nn by (auto intro!: nn_integral_cong_AE elim!: eventually_elim2)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1841
      also note fin
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1842
      finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1843
        by simp
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1844
    qed simp
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1845
  qed fact
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1846
  finally show ?thesis
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1847
    using nn by (simp add: integral_nonneg_AE)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1848
qed
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1849
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1850
lemma has_bochner_integral_nn_integral:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1851
  assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1852
  assumes "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1853
  shows "has_bochner_integral M f x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1854
  unfolding has_bochner_integral_iff
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1855
  using assms by (auto simp: assms integral_eq_nn_integral intro: integrableI_nn_integral_finite)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1856
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1857
lemma integrableI_simple_bochner_integrable:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1858
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1859
  shows "simple_bochner_integrable M f \<Longrightarrow> integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1860
  by (intro integrableI_sequence[where s="\<lambda>_. f"] borel_measurable_simple_function)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1861
     (auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1862
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1863
proposition integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1864
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1865
  assumes "integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1866
  assumes base: "\<And>A c. A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> P (\<lambda>x. indicator A x *\<^sub>R c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1867
  assumes add: "\<And>f g. integrable M f \<Longrightarrow> P f \<Longrightarrow> integrable M g \<Longrightarrow> P g \<Longrightarrow> P (\<lambda>x. f x + g x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1868
  assumes lim: "\<And>f s. (\<And>i. integrable M (s i)) \<Longrightarrow> (\<And>i. P (s i)) \<Longrightarrow>
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1869
   (\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x) \<Longrightarrow>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1870
   (\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)) \<Longrightarrow> integrable M f \<Longrightarrow> P f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1871
  shows "P f"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1872
proof -
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
  1873
  from \<open>integrable M f\<close> have f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1874
    unfolding integrable_iff_bounded by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1875
  from borel_measurable_implies_sequence_metric[OF f(1)]
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1876
  obtain s where s: "\<And>i. simple_function M (s i)" "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1877
    "\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1878
    unfolding norm_conv_dist by metis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1879
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1880
  { fix f A
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1881
    have [simp]: "P (\<lambda>x. 0)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1882
      using base[of "{}" undefined] by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1883
    have "(\<And>i::'b. i \<in> A \<Longrightarrow> integrable M (f i::'a \<Rightarrow> 'b)) \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1884
    (\<And>i. i \<in> A \<Longrightarrow> P (f i)) \<Longrightarrow> P (\<lambda>x. \<Sum>i\<in>A. f i x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1885
    by (induct A rule: infinite_finite_induct) (auto intro!: add) }
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1886
  note sum = this
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1887
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
  1888
  define s' where [abs_def]: "s' i z = indicator (space M) z *\<^sub>R s i z" for i z
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1889
  then have s'_eq_s: "\<And>i x. x \<in> space M \<Longrightarrow> s' i x = s i x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1890
    by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1891
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1892
  have sf[measurable]: "\<And>i. simple_function M (s' i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1893
    unfolding s'_def using s(1)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1894
    by (intro simple_function_compose2[where h="(*\<^sub>R)"] simple_function_indicator) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1895
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1896
  { fix i
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1897
    have "\<And>z. {y. s' i z = y \<and> y \<in> s' i ` space M \<and> y \<noteq> 0 \<and> z \<in> space M} =
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1898
        (if z \<in> space M \<and> s' i z \<noteq> 0 then {s' i z} else {})"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1899
      by (auto simp add: s'_def split: split_indicator)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1900
    then have "\<And>z. s' i = (\<lambda>z. \<Sum>y\<in>s' i`space M - {0}. indicator {x\<in>space M. s' i x = y} z *\<^sub>R y)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1901
      using sf by (auto simp: fun_eq_iff simple_function_def s'_def) }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1902
  note s'_eq = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1903
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1904
  show "P f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1905
  proof (rule lim)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1906
    fix i
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1907
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1908
    have "(\<integral>\<^sup>+x. norm (s' i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1909
      using s by (intro nn_integral_mono) (auto simp: s'_eq_s)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1910
    also have "\<dots> < \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1911
      using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1912
    finally have sbi: "simple_bochner_integrable M (s' i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1913
      using sf by (intro simple_bochner_integrableI_bounded) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1914
    then show "integrable M (s' i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1915
      by (rule integrableI_simple_bochner_integrable)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1916
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1917
    { fix x assume"x \<in> space M" "s' i x \<noteq> 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1918
      then have "emeasure M {y \<in> space M. s' i y = s' i x} \<le> emeasure M {y \<in> space M. s' i y \<noteq> 0}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1919
        by (intro emeasure_mono) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1920
      also have "\<dots> < \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1921
        using sbi by (auto elim: simple_bochner_integrable.cases simp: less_top)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1922
      finally have "emeasure M {y \<in> space M. s' i y = s' i x} \<noteq> \<infinity>" by simp }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1923
    then show "P (s' i)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1924
      by (subst s'_eq) (auto intro!: sum base simp: less_top)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1925
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1926
    fix x assume "x \<in> space M" with s show "(\<lambda>i. s' i x) \<longlonglongrightarrow> f x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1927
      by (simp add: s'_eq_s)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1928
    show "norm (s' i x) \<le> 2 * norm (f x)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
  1929
      using \<open>x \<in> space M\<close> s by (simp add: s'_eq_s)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1930
  qed fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1931
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1932
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1933
lemma integral_eq_zero_AE:
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1934
  "(AE x in M. f x = 0) \<Longrightarrow> integral\<^sup>L M f = 0"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1935
  using integral_cong_AE[of f M "\<lambda>_. 0"]
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  1936
  by (cases "integrable M f") (simp_all add: not_integrable_integral_eq)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1937
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1938
lemma integral_nonneg_eq_0_iff_AE:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1939
  fixes f :: "_ \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1940
  assumes f[measurable]: "integrable M f" and nonneg: "AE x in M. 0 \<le> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1941
  shows "integral\<^sup>L M f = 0 \<longleftrightarrow> (AE x in M. f x = 0)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1942
proof
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1943
  assume "integral\<^sup>L M f = 0"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1944
  then have "integral\<^sup>N M f = 0"
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1945
    using nn_integral_eq_integral[OF f nonneg] by simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1946
  then have "AE x in M. ennreal (f x) \<le> 0"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1947
    by (simp add: nn_integral_0_iff_AE)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1948
  with nonneg show "AE x in M. f x = 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1949
    by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1950
qed (auto simp add: integral_eq_zero_AE)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1951
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1952
lemma integral_mono_AE:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1953
  fixes f :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1954
  assumes "integrable M f" "integrable M g" "AE x in M. f x \<le> g x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1955
  shows "integral\<^sup>L M f \<le> integral\<^sup>L M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1956
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1957
  have "0 \<le> integral\<^sup>L M (\<lambda>x. g x - f x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1958
    using assms by (intro integral_nonneg_AE integrable_diff assms) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1959
  also have "\<dots> = integral\<^sup>L M g - integral\<^sup>L M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1960
    by (intro integral_diff assms)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1961
  finally show ?thesis by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1962
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1963
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1964
lemma integral_mono:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1965
  fixes f :: "'a \<Rightarrow> real"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1966
  shows "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x \<le> g x) \<Longrightarrow>
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1967
    integral\<^sup>L M f \<le> integral\<^sup>L M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1968
  by (intro integral_mono_AE) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  1969
70271
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1970
lemma integral_norm_bound_integral:
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1971
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{banach,second_countable_topology}"
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1972
  assumes "integrable M f" "integrable M g" "\<And>x. x \<in> space M \<Longrightarrow> norm(f x) \<le> g x"
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1973
  shows "norm (\<integral>x. f x \<partial>M) \<le> (\<integral>x. g x \<partial>M)"
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1974
proof -
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1975
  have "norm (\<integral>x. f x \<partial>M) \<le> (\<integral>x. norm (f x) \<partial>M)"
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1976
    by (rule integral_norm_bound)
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1977
  also have "... \<le> (\<integral>x. g x \<partial>M)"
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1978
    using assms integrable_norm integral_mono by blast
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1979
  finally show ?thesis .
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1980
qed
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1981
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1982
lemma integral_abs_bound_integral:
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1983
  fixes f :: "'a::euclidean_space \<Rightarrow> real"
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1984
  assumes "integrable M f" "integrable M g" "\<And>x. x \<in> space M \<Longrightarrow> \<bar>f x\<bar> \<le> g x"
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1985
  shows "\<bar>\<integral>x. f x \<partial>M\<bar> \<le> (\<integral>x. g x \<partial>M)"
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1986
  by (metis integral_norm_bound_integral assms real_norm_def)
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1987
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64287
diff changeset
  1988
text \<open>The next two statements are useful to bound Lebesgue integrals, as they avoid one
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1989
integrability assumption. The price to pay is that the upper function has to be nonnegative,
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64287
diff changeset
  1990
but this is often true and easy to check in computations.\<close>
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1991
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1992
lemma integral_mono_AE':
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1993
  fixes f::"_ \<Rightarrow> real"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1994
  assumes "integrable M f" "AE x in M. g x \<le> f x" "AE x in M. 0 \<le> f x"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1995
  shows "(\<integral>x. g x \<partial>M) \<le> (\<integral>x. f x \<partial>M)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  1996
proof (cases "integrable M g")
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1997
  case True
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1998
  show ?thesis by (rule integral_mono_AE, auto simp add: assms True)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  1999
next
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2000
  case False
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2001
  then have "(\<integral>x. g x \<partial>M) = 0" by (simp add: not_integrable_integral_eq)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2002
  also have "... \<le> (\<integral>x. f x \<partial>M)" by (simp add: integral_nonneg_AE[OF assms(3)])
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2003
  finally show ?thesis by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2004
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2005
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2006
lemma integral_mono':
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2007
  fixes f::"_ \<Rightarrow> real"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2008
  assumes "integrable M f" "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2009
  shows "(\<integral>x. g x \<partial>M) \<le> (\<integral>x. f x \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2010
by (rule integral_mono_AE', insert assms, auto)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2011
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2012
lemma (in finite_measure) integrable_measure:
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2013
  assumes I: "disjoint_family_on X I" "countable I"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2014
  shows "integrable (count_space I) (\<lambda>i. measure M (X i))"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2015
proof -
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2016
  have "(\<integral>\<^sup>+i. measure M (X i) \<partial>count_space I) = (\<integral>\<^sup>+i. measure M (if X i \<in> sets M then X i else {}) \<partial>count_space I)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2017
    by (auto intro!: nn_integral_cong measure_notin_sets)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2018
  also have "\<dots> = measure M (\<Union>i\<in>I. if X i \<in> sets M then X i else {})"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2019
    using I unfolding emeasure_eq_measure[symmetric]
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2020
    by (subst emeasure_UN_countable) (auto simp: disjoint_family_on_def)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2021
  finally show ?thesis
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2022
    by (auto intro!: integrableI_bounded)
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2023
qed
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2024
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2025
lemma integrableI_real_bounded:
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2026
  assumes f: "f \<in> borel_measurable M" and ae: "AE x in M. 0 \<le> f x" and fin: "integral\<^sup>N M f < \<infinity>"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2027
  shows "integrable M f"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2028
proof (rule integrableI_bounded)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2029
  have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = \<integral>\<^sup>+ x. ennreal (f x) \<partial>M"
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2030
    using ae by (auto intro: nn_integral_cong_AE)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2031
  also note fin
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2032
  finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2033
qed fact
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2034
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2035
lemma nn_integral_nonneg_infinite:
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2036
  fixes f::"'a \<Rightarrow> real"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2037
  assumes "f \<in> borel_measurable M" "\<not> integrable M f" "AE x in M. f x \<ge> 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2038
  shows "(\<integral>\<^sup>+x. f x \<partial>M) = \<infinity>"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2039
using assms integrableI_real_bounded less_top by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2040
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2041
lemma integral_real_bounded:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2042
  assumes "0 \<le> r" "integral\<^sup>N M f \<le> ennreal r"
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2043
  shows "integral\<^sup>L M f \<le> r"
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2044
proof cases
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2045
  assume [simp]: "integrable M f"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2046
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2047
  have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = integral\<^sup>N M (\<lambda>x. max 0 (f x))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2048
    by (intro nn_integral_eq_integral[symmetric]) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2049
  also have "\<dots> = integral\<^sup>N M f"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2050
    by (intro nn_integral_cong) (simp add: max_def ennreal_neg)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2051
  also have "\<dots> \<le> r"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2052
    by fact
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2053
  finally have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) \<le> r"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2054
    using \<open>0 \<le> r\<close> by simp
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2055
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2056
  moreover have "integral\<^sup>L M f \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2057
    by (rule integral_mono_AE) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2058
  ultimately show ?thesis
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2059
    by simp
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2060
next
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2061
  assume "\<not> integrable M f" then show ?thesis
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2062
    using \<open>0 \<le> r\<close> by (simp add: not_integrable_integral_eq)
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2063
qed
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
  2064
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2065
lemma integrable_MIN:
68794
63e84bd8e1f6 moved lemma from AFP
nipkow
parents: 68403
diff changeset
  2066
  fixes f:: "_ \<Rightarrow> _ \<Rightarrow> real"
68798
07714b60f653 tuned proofs
nipkow
parents: 68794
diff changeset
  2067
  shows "\<lbrakk> finite I;  I \<noteq> {}; \<And>i. i \<in> I \<Longrightarrow> integrable M (f i) \<rbrakk>
07714b60f653 tuned proofs
nipkow
parents: 68794
diff changeset
  2068
   \<Longrightarrow> integrable M (\<lambda>x. MIN i\<in>I. f i x)"
07714b60f653 tuned proofs
nipkow
parents: 68794
diff changeset
  2069
by (induct rule: finite_ne_induct) simp+
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2070
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2071
lemma integrable_MAX:
68794
63e84bd8e1f6 moved lemma from AFP
nipkow
parents: 68403
diff changeset
  2072
  fixes f:: "_ \<Rightarrow> _ \<Rightarrow> real"
68798
07714b60f653 tuned proofs
nipkow
parents: 68794
diff changeset
  2073
  shows "\<lbrakk> finite I;  I \<noteq> {};  \<And>i. i \<in> I \<Longrightarrow> integrable M (f i) \<rbrakk>
07714b60f653 tuned proofs
nipkow
parents: 68794
diff changeset
  2074
  \<Longrightarrow> integrable M (\<lambda>x. MAX i\<in>I. f i x)"
07714b60f653 tuned proofs
nipkow
parents: 68794
diff changeset
  2075
by (induct rule: finite_ne_induct) simp+
68794
63e84bd8e1f6 moved lemma from AFP
nipkow
parents: 68403
diff changeset
  2076
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2077
theorem integral_Markov_inequality:
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2078
  assumes [measurable]: "integrable M u" and "AE x in M. 0 \<le> u x" "0 < (c::real)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2079
  shows "(emeasure M) {x\<in>space M. u x \<ge> c} \<le> (1/c) * (\<integral>x. u x \<partial>M)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2080
proof -
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2081
  have "(\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M) \<le> (\<integral>\<^sup>+ x. u x \<partial>M)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64287
diff changeset
  2082
    by (rule nn_integral_mono_AE, auto simp add: \<open>c>0\<close> less_eq_real_def)
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2083
  also have "... = (\<integral>x. u x \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2084
    by (rule nn_integral_eq_integral, auto simp add: assms)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2085
  finally have *: "(\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M) \<le> (\<integral>x. u x \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2086
    by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2087
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2088
  have "{x \<in> space M. u x \<ge> c} = {x \<in> space M. ennreal(1/c) * u x \<ge> 1} \<inter> (space M)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64287
diff changeset
  2089
    using \<open>c>0\<close> by (auto simp: ennreal_mult'[symmetric])
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2090
  then have "emeasure M {x \<in> space M. u x \<ge> c} = emeasure M ({x \<in> space M. ennreal(1/c) * u x \<ge> 1})"
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2091
    by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2092
  also have "... \<le> ennreal(1/c) * (\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2093
    by (rule nn_integral_Markov_inequality) (auto simp add: assms)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2094
  also have "... \<le> ennreal(1/c) * (\<integral>x. u x \<partial>M)"
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2095
    by (rule mult_left_mono) (use * \<open>c > 0\<close> in auto)
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2096
  finally show ?thesis
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2097
    using \<open>0<c\<close> by (simp add: ennreal_mult'[symmetric])
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2098
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2099
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2100
theorem integral_Markov_inequality_measure:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2101
  assumes [measurable]: "integrable M u" and "A \<in> sets M" and "AE x in M. 0 \<le> u x" "0 < (c::real)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2102
  shows "measure M {x\<in>space M. u x \<ge> c} \<le> (\<integral>x. u x \<partial>M) / c"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2103
proof -
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2104
  have le: "emeasure M {x\<in>space M. u x \<ge> c} \<le> ennreal ((1/c) * (\<integral>x. u x \<partial>M))"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2105
    by (rule integral_Markov_inequality) (use assms in auto)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2106
  also have "\<dots> < top"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2107
    by auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2108
  finally have "ennreal (measure M {x\<in>space M. u x \<ge> c}) = emeasure M {x\<in>space M. u x \<ge> c}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2109
    by (intro emeasure_eq_ennreal_measure [symmetric]) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2110
  also note le
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2111
  finally show ?thesis
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2112
    by (subst (asm) ennreal_le_iff)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2113
       (auto intro!: divide_nonneg_pos Bochner_Integration.integral_nonneg_AE assms)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2114
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2115
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2116
theorem%important (in finite_measure) second_moment_method:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2117
  assumes [measurable]: "f \<in> M \<rightarrow>\<^sub>M borel"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2118
  assumes "integrable M (\<lambda>x. f x ^ 2)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2119
  defines "\<mu> \<equiv> lebesgue_integral M f"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2120
  assumes "a > 0"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2121
  shows "measure M {x\<in>space M. \<bar>f x\<bar> \<ge> a} \<le> lebesgue_integral M (\<lambda>x. f x ^ 2) / a\<^sup>2"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2122
proof -
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2123
  have integrable: "integrable M f"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2124
    using assms by (blast dest: square_integrable_imp_integrable)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2125
  have "{x\<in>space M. \<bar>f x\<bar> \<ge> a} = {x\<in>space M. f x ^ 2 \<ge> a\<^sup>2}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2126
    using \<open>a > 0\<close> by (simp flip: abs_le_square_iff)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2127
  hence "measure M {x\<in>space M. \<bar>f x\<bar> \<ge> a} = measure M {x\<in>space M. f x ^ 2 \<ge> a\<^sup>2}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2128
    by simp
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2129
  also have "\<dots> \<le> lebesgue_integral M (\<lambda>x. f x ^ 2) / a\<^sup>2"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2130
    using assms by (intro integral_Markov_inequality_measure) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2131
  finally show ?thesis .
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2132
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 71840
diff changeset
  2133
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2134
lemma integral_ineq_eq_0_then_AE:
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2135
  fixes f::"_ \<Rightarrow> real"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2136
  assumes "AE x in M. f x \<le> g x" "integrable M f" "integrable M g"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2137
          "(\<integral>x. f x \<partial>M) = (\<integral>x. g x \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2138
  shows "AE x in M. f x = g x"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2139
proof -
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2140
  define h where "h = (\<lambda>x. g x - f x)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2141
  have "AE x in M. h x = 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2142
    apply (subst integral_nonneg_eq_0_iff_AE[symmetric])
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2143
    unfolding h_def using assms by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2144
  then show ?thesis unfolding h_def by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2145
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2146
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2147
lemma not_AE_zero_int_E:
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2148
  fixes f::"'a \<Rightarrow> real"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2149
  assumes "AE x in M. f x \<ge> 0" "(\<integral>x. f x \<partial>M) > 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2150
      and [measurable]: "f \<in> borel_measurable M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2151
  shows "\<exists>A e. A \<in> sets M \<and> e>0 \<and> emeasure M A > 0 \<and> (\<forall>x \<in> A. f x \<ge> e)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2152
proof (rule not_AE_zero_E, auto simp add: assms)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2153
  assume *: "AE x in M. f x = 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2154
  have "(\<integral>x. f x \<partial>M) = (\<integral>x. 0 \<partial>M)" by (rule integral_cong_AE, auto simp add: *)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2155
  then have "(\<integral>x. f x \<partial>M) = 0" by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2156
  then show False using assms(2) by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2157
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2158
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2159
proposition tendsto_L1_int:
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2160
  fixes u :: "_ \<Rightarrow> _ \<Rightarrow> 'b::{banach, second_countable_topology}"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2161
  assumes [measurable]: "\<And>n. integrable M (u n)" "integrable M f"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2162
          and "((\<lambda>n. (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)) \<longlongrightarrow> 0) F"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2163
  shows "((\<lambda>n. (\<integral>x. u n x \<partial>M)) \<longlongrightarrow> (\<integral>x. f x \<partial>M)) F"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2164
proof -
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2165
  have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> (0::ennreal)) F"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2166
  proof (rule tendsto_sandwich[of "\<lambda>_. 0", where ?h = "\<lambda>n. (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)"], auto simp add: assms)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2167
    {
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2168
      fix n
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2169
      have "(\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M) = (\<integral>x. u n x - f x \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2170
        apply (rule Bochner_Integration.integral_diff[symmetric]) using assms by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2171
      then have "norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) = norm (\<integral>x. u n x - f x \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2172
        by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2173
      also have "... \<le> (\<integral>x. norm(u n x - f x) \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2174
        by (rule integral_norm_bound)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2175
      finally have "ennreal(norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<le> (\<integral>x. norm(u n x - f x) \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2176
        by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2177
      also have "... = (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2178
        apply (rule nn_integral_eq_integral[symmetric]) using assms by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2179
      finally have "norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)" by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2180
    }
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2181
    then show "eventually (\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)) F"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2182
      by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2183
  qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2184
  then have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F"
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68073
diff changeset
  2185
    by (simp flip: ennreal_0)
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2186
  then have "((\<lambda>n. ((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F" using tendsto_norm_zero_iff by blast
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2187
  then show ?thesis using Lim_null by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2188
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2189
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69546
diff changeset
  2190
text \<open>The next lemma asserts that, if a sequence of functions converges in \<open>L\<^sup>1\<close>, then
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64287
diff changeset
  2191
it admits a subsequence that converges almost everywhere.\<close>
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2192
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2193
proposition tendsto_L1_AE_subseq:
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2194
  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2195
  assumes [measurable]: "\<And>n. integrable M (u n)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2196
      and "(\<lambda>n. (\<integral>x. norm(u n x) \<partial>M)) \<longlonglongrightarrow> 0"
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 65680
diff changeset
  2197
  shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (AE x in M. (\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2198
proof -
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2199
  {
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2200
    fix k
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2201
    have "eventually (\<lambda>n. (\<integral>x. norm(u n x) \<partial>M) < (1/2)^k) sequentially"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2202
      using order_tendstoD(2)[OF assms(2)] by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2203
    with eventually_elim2[OF eventually_gt_at_top[of k] this]
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2204
    have "\<exists>n>k. (\<integral>x. norm(u n x) \<partial>M) < (1/2)^k"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2205
      by (metis eventually_False_sequentially)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2206
  }
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2207
  then have "\<exists>r. \<forall>n. True \<and> (r (Suc n) > r n \<and> (\<integral>x. norm(u (r (Suc n)) x) \<partial>M) < (1/2)^(r n))"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2208
    by (intro dependent_nat_choice, auto)
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 65680
diff changeset
  2209
  then obtain r0 where r0: "strict_mono r0" "\<And>n. (\<integral>x. norm(u (r0 (Suc n)) x) \<partial>M) < (1/2)^(r0 n)"
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 65680
diff changeset
  2210
    by (auto simp: strict_mono_Suc_iff)
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2211
  define r where "r = (\<lambda>n. r0(n+1))"
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 65680
diff changeset
  2212
  have "strict_mono r" unfolding r_def using r0(1) by (simp add: strict_mono_Suc_iff)
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2213
  have I: "(\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M) < ennreal((1/2)^n)" for n
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2214
  proof -
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 65680
diff changeset
  2215
    have "r0 n \<ge> n" using \<open>strict_mono r0\<close> by (simp add: seq_suble)
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64287
diff changeset
  2216
    have "(1/2::real)^(r0 n) \<le> (1/2)^n" by (rule power_decreasing[OF \<open>r0 n \<ge> n\<close>], auto)
69700
7a92cbec7030 new material about summations and powers, along with some tweaks
paulson <lp15@cam.ac.uk>
parents: 69683
diff changeset
  2217
    then have "(\<integral>x. norm(u (r0 (Suc n)) x) \<partial>M) < (1/2)^n" 
7a92cbec7030 new material about summations and powers, along with some tweaks
paulson <lp15@cam.ac.uk>
parents: 69683
diff changeset
  2218
      using r0(2) less_le_trans by blast
7a92cbec7030 new material about summations and powers, along with some tweaks
paulson <lp15@cam.ac.uk>
parents: 69683
diff changeset
  2219
    then have "(\<integral>x. norm(u (r n) x) \<partial>M) < (1/2)^n" 
7a92cbec7030 new material about summations and powers, along with some tweaks
paulson <lp15@cam.ac.uk>
parents: 69683
diff changeset
  2220
      unfolding r_def by auto
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2221
    moreover have "(\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M) = (\<integral>x. norm(u (r n) x) \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2222
      by (rule nn_integral_eq_integral, auto simp add: integrable_norm[OF assms(1)[of "r n"]])
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2223
    ultimately show ?thesis by (auto intro: ennreal_lessI)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2224
  qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2225
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2226
  have "AE x in M. limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2227
  proof (rule AE_upper_bound_inf_ennreal)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2228
    fix e::real assume "e > 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2229
    define A where "A = (\<lambda>n. {x \<in> space M. norm(u (r n) x) \<ge> e})"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2230
    have A_meas [measurable]: "\<And>n. A n \<in> sets M" unfolding A_def by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2231
    have A_bound: "emeasure M (A n) < (1/e) * ennreal((1/2)^n)" for n
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2232
    proof -
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2233
      have *: "indicator (A n) x \<le> (1/e) * ennreal(norm(u (r n) x))" for x
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2234
        apply (cases "x \<in> A n") unfolding A_def using \<open>0 < e\<close> by (auto simp: ennreal_mult[symmetric])
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2235
      have "emeasure M (A n) = (\<integral>\<^sup>+x. indicator (A n) x \<partial>M)" by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2236
      also have "... \<le> (\<integral>\<^sup>+x. (1/e) * ennreal(norm(u (r n) x)) \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2237
        apply (rule nn_integral_mono) using * by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2238
      also have "... = (1/e) * (\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64287
diff changeset
  2239
        apply (rule nn_integral_cmult) using \<open>e > 0\<close> by auto
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2240
      also have "... < (1/e) * ennreal((1/2)^n)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64287
diff changeset
  2241
        using I[of n] \<open>e > 0\<close> by (intro ennreal_mult_strict_left_mono) auto
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2242
      finally show ?thesis by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2243
    qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2244
    have A_fin: "emeasure M (A n) < \<infinity>" for n
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64287
diff changeset
  2245
      using \<open>e > 0\<close> A_bound[of n]
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2246
      by (auto simp add: ennreal_mult less_top[symmetric])
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2247
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2248
    have A_sum: "summable (\<lambda>n. measure M (A n))"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2249
    proof (rule summable_comparison_test'[of "\<lambda>n. (1/e) * (1/2)^n" 0])
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2250
      have "summable (\<lambda>n. (1/(2::real))^n)" by (simp add: summable_geometric)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2251
      then show "summable (\<lambda>n. (1/e) * (1/2)^n)" using summable_mult by blast
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2252
      fix n::nat assume "n \<ge> 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2253
      have "norm(measure M (A n)) = measure M (A n)" by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2254
      also have "... = enn2real(emeasure M (A n))" unfolding measure_def by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2255
      also have "... < enn2real((1/e) * (1/2)^n)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2256
        using A_bound[of n] \<open>emeasure M (A n) < \<infinity>\<close> \<open>0 < e\<close>
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2257
        by (auto simp: emeasure_eq_ennreal_measure ennreal_mult[symmetric] ennreal_less_iff)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2258
      also have "... = (1/e) * (1/2)^n"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2259
        using \<open>0<e\<close> by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2260
      finally show "norm(measure M (A n)) \<le> (1/e) * (1/2)^n" by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2261
    qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2262
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2263
    have "AE x in M. eventually (\<lambda>n. x \<in> space M - A n) sequentially"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2264
      by (rule borel_cantelli_AE1[OF A_meas A_fin A_sum])
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2265
    moreover
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2266
    {
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2267
      fix x assume "eventually (\<lambda>n. x \<in> space M - A n) sequentially"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2268
      moreover have "norm(u (r n) x) \<le> ennreal e" if "x \<in> space M - A n" for n
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2269
        using that unfolding A_def by (auto intro: ennreal_leI)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2270
      ultimately have "eventually (\<lambda>n. norm(u (r n) x) \<le> ennreal e) sequentially"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2271
        by (simp add: eventually_mono)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2272
      then have "limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> e"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2273
        by (simp add: Limsup_bounded)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2274
    }
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2275
    ultimately show "AE x in M. limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0 + ennreal e" by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2276
  qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2277
  moreover
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2278
  {
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2279
    fix x assume "limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2280
    moreover then have "liminf (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2281
      by (rule order_trans[rotated]) (auto intro: Liminf_le_Limsup)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2282
    ultimately have "(\<lambda>n. ennreal (norm(u (r n) x))) \<longlonglongrightarrow> 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2283
      using tendsto_Limsup[of sequentially "\<lambda>n. ennreal (norm(u (r n) x))"] by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2284
    then have "(\<lambda>n. norm(u (r n) x)) \<longlonglongrightarrow> 0"
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68073
diff changeset
  2285
      by (simp flip: ennreal_0)
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2286
    then have "(\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2287
      by (simp add: tendsto_norm_zero_iff)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2288
  }
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2289
  ultimately have "AE x in M. (\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0" by auto
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 65680
diff changeset
  2290
  then show ?thesis using \<open>strict_mono r\<close> by auto
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2291
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2292
69683
8b3458ca0762 subsection is always %important
immler
parents: 69652
diff changeset
  2293
subsection \<open>Restricted measure spaces\<close>
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68798
diff changeset
  2294
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2295
lemma integrable_restrict_space:
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2296
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2297
  assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2298
  shows "integrable (restrict_space M \<Omega>) f \<longleftrightarrow> integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2299
  unfolding integrable_iff_bounded
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2300
    borel_measurable_restrict_space_iff[OF \<Omega>]
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2301
    nn_integral_restrict_space[OF \<Omega>]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2302
  by (simp add: ac_simps ennreal_indicator ennreal_mult)
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2303
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2304
lemma integral_restrict_space:
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2305
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2306
  assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2307
  shows "integral\<^sup>L (restrict_space M \<Omega>) f = integral\<^sup>L M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2308
proof (rule integral_eq_cases)
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2309
  assume "integrable (restrict_space M \<Omega>) f"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2310
  then show ?thesis
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2311
  proof induct
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2312
    case (base A c) then show ?case
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2313
      by (simp add: indicator_inter_arith[symmetric] sets_restrict_space_iff
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2314
                    emeasure_restrict_space Int_absorb1 measure_restrict_space)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2315
  next
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2316
    case (add g f) then show ?case
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2317
      by (simp add: scaleR_add_right integrable_restrict_space)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2318
  next
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2319
    case (lim f s)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2320
    show ?case
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2321
    proof (rule LIMSEQ_unique)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2322
      show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) \<longlonglongrightarrow> integral\<^sup>L (restrict_space M \<Omega>) f"
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2323
        using lim by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) simp_all
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2324
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2325
      show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) \<longlonglongrightarrow> (\<integral> x. indicator \<Omega> x *\<^sub>R f x \<partial>M)"
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2326
        unfolding lim
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2327
        using lim
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2328
        by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (indicator \<Omega> x *\<^sub>R f x)"])
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2329
           (auto simp add: space_restrict_space integrable_restrict_space simp del: norm_scaleR
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2330
                 split: split_indicator)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2331
    qed
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2332
  qed
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2333
qed (simp add: integrable_restrict_space)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2334
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2335
lemma integral_empty:
60066
14efa7f4ee7b add lemmas
Andreas Lochbihler
parents: 59867
diff changeset
  2336
  assumes "space M = {}"
14efa7f4ee7b add lemmas
Andreas Lochbihler
parents: 59867
diff changeset
  2337
  shows "integral\<^sup>L M f = 0"
14efa7f4ee7b add lemmas
Andreas Lochbihler
parents: 59867
diff changeset
  2338
proof -
14efa7f4ee7b add lemmas
Andreas Lochbihler
parents: 59867
diff changeset
  2339
  have "(\<integral> x. f x \<partial>M) = (\<integral> x. 0 \<partial>M)"
14efa7f4ee7b add lemmas
Andreas Lochbihler
parents: 59867
diff changeset
  2340
    by(rule integral_cong)(simp_all add: assms)
14efa7f4ee7b add lemmas
Andreas Lochbihler
parents: 59867
diff changeset
  2341
  thus ?thesis by simp
14efa7f4ee7b add lemmas
Andreas Lochbihler
parents: 59867
diff changeset
  2342
qed
14efa7f4ee7b add lemmas
Andreas Lochbihler
parents: 59867
diff changeset
  2343
69683
8b3458ca0762 subsection is always %important
immler
parents: 69652
diff changeset
  2344
subsection \<open>Measure spaces with an associated density\<close>
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68798
diff changeset
  2345
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2346
lemma integrable_density:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2347
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2348
  assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2349
    and nn: "AE x in M. 0 \<le> g x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2350
  shows "integrable (density M g) f \<longleftrightarrow> integrable M (\<lambda>x. g x *\<^sub>R f x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2351
  unfolding integrable_iff_bounded using nn
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2352
  apply (simp add: nn_integral_density less_top[symmetric])
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66804
diff changeset
  2353
  apply (intro arg_cong2[where f="(=)"] refl nn_integral_cong_AE)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2354
  apply (auto simp: ennreal_mult)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2355
  done
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2356
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2357
lemma integral_density:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2358
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2359
  assumes f: "f \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2360
    and g[measurable]: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2361
  shows "integral\<^sup>L (density M g) f = integral\<^sup>L M (\<lambda>x. g x *\<^sub>R f x)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2362
proof (rule integral_eq_cases)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2363
  assume "integrable (density M g) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2364
  then show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2365
  proof induct
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2366
    case (base A c)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2367
    then have [measurable]: "A \<in> sets M" by auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2368
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2369
    have int: "integrable M (\<lambda>x. g x * indicator A x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2370
      using g base integrable_density[of "indicator A :: 'a \<Rightarrow> real" M g] by simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2371
    then have "integral\<^sup>L M (\<lambda>x. g x * indicator A x) = (\<integral>\<^sup>+ x. ennreal (g x * indicator A x) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2372
      using g by (subst nn_integral_eq_integral) auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2373
    also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (g x) * indicator A x \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2374
      by (intro nn_integral_cong) (auto split: split_indicator)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2375
    also have "\<dots> = emeasure (density M g) A"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2376
      by (rule emeasure_density[symmetric]) auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2377
    also have "\<dots> = ennreal (measure (density M g) A)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2378
      using base by (auto intro: emeasure_eq_ennreal_measure)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2379
    also have "\<dots> = integral\<^sup>L (density M g) (indicator A)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2380
      using base by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2381
    finally show ?case
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2382
      using base g
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2383
      apply (simp add: int integral_nonneg_AE)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2384
      apply (subst (asm) ennreal_inj)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2385
      apply (auto intro!: integral_nonneg_AE)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2386
      done
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2387
  next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2388
    case (add f h)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2389
    then have [measurable]: "f \<in> borel_measurable M" "h \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2390
      by (auto dest!: borel_measurable_integrable)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2391
    from add g show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2392
      by (simp add: scaleR_add_right integrable_density)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2393
  next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2394
    case (lim f s)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2395
    have [measurable]: "f \<in> borel_measurable M" "\<And>i. s i \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2396
      using lim(1,5)[THEN borel_measurable_integrable] by auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2397
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2398
    show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2399
    proof (rule LIMSEQ_unique)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2400
      show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. g x *\<^sub>R f x)"
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2401
      proof (rule integral_dominated_convergence)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2402
        show "integrable M (\<lambda>x. 2 * norm (g x *\<^sub>R f x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2403
          by (intro integrable_mult_right integrable_norm integrable_density[THEN iffD1] lim g) auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2404
        show "AE x in M. (\<lambda>i. g x *\<^sub>R s i x) \<longlonglongrightarrow> g x *\<^sub>R f x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2405
          using lim(3) by (auto intro!: tendsto_scaleR AE_I2[of M])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2406
        show "\<And>i. AE x in M. norm (g x *\<^sub>R s i x) \<le> 2 * norm (g x *\<^sub>R f x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2407
          using lim(4) g by (auto intro!: AE_I2[of M] mult_left_mono simp: field_simps)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2408
      qed auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2409
      show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) \<longlonglongrightarrow> integral\<^sup>L (density M g) f"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2410
        unfolding lim(2)[symmetric]
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2411
        by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  2412
           (insert lim(3-5), auto)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2413
    qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2414
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2415
qed (simp add: f g integrable_density)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2416
69739
nipkow
parents: 69722
diff changeset
  2417
lemma
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2418
  fixes g :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2419
  assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "g \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2420
  shows integral_real_density: "integral\<^sup>L (density M f) g = (\<integral> x. f x * g x \<partial>M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2421
    and integrable_real_density: "integrable (density M f) g \<longleftrightarrow> integrable M (\<lambda>x. f x * g x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2422
  using assms integral_density[of g M f] integrable_density[of g M f] by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2423
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2424
lemma has_bochner_integral_density:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2425
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2426
  shows "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. 0 \<le> g x) \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2427
    has_bochner_integral M (\<lambda>x. g x *\<^sub>R f x) x \<Longrightarrow> has_bochner_integral (density M g) f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2428
  by (simp add: has_bochner_integral_iff integrable_density integral_density)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2429
69683
8b3458ca0762 subsection is always %important
immler
parents: 69652
diff changeset
  2430
subsection \<open>Distributions\<close>
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68798
diff changeset
  2431
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2432
lemma integrable_distr_eq:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2433
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2434
  assumes [measurable]: "g \<in> measurable M N" "f \<in> borel_measurable N"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2435
  shows "integrable (distr M N g) f \<longleftrightarrow> integrable M (\<lambda>x. f (g x))"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2436
  unfolding integrable_iff_bounded by (simp_all add: nn_integral_distr)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2437
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2438
lemma integrable_distr:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2439
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2440
  shows "T \<in> measurable M M' \<Longrightarrow> integrable (distr M M' T) f \<Longrightarrow> integrable M (\<lambda>x. f (T x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2441
  by (subst integrable_distr_eq[symmetric, where g=T])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2442
     (auto dest: borel_measurable_integrable)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2443
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2444
lemma integral_distr:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2445
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2446
  assumes g[measurable]: "g \<in> measurable M N" and f: "f \<in> borel_measurable N"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2447
  shows "integral\<^sup>L (distr M N g) f = integral\<^sup>L M (\<lambda>x. f (g x))"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2448
proof (rule integral_eq_cases)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2449
  assume "integrable (distr M N g) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2450
  then show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2451
  proof induct
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2452
    case (base A c)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2453
    then have [measurable]: "A \<in> sets N" by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2454
    from base have int: "integrable (distr M N g) (\<lambda>a. indicator A a *\<^sub>R c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2455
      by (intro integrable_indicator)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2456
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2457
    have "integral\<^sup>L (distr M N g) (\<lambda>a. indicator A a *\<^sub>R c) = measure (distr M N g) A *\<^sub>R c"
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2458
      using base by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2459
    also have "\<dots> = measure M (g -` A \<inter> space M) *\<^sub>R c"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2460
      by (subst measure_distr) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2461
    also have "\<dots> = integral\<^sup>L M (\<lambda>a. indicator (g -` A \<inter> space M) a *\<^sub>R c)"
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2462
      using base by (auto simp: emeasure_distr)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2463
    also have "\<dots> = integral\<^sup>L M (\<lambda>a. indicator A (g a) *\<^sub>R c)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2464
      using int base by (intro integral_cong_AE) (auto simp: emeasure_distr split: split_indicator)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2465
    finally show ?case .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2466
  next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2467
    case (add f h)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2468
    then have [measurable]: "f \<in> borel_measurable N" "h \<in> borel_measurable N"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2469
      by (auto dest!: borel_measurable_integrable)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2470
    from add g show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2471
      by (simp add: scaleR_add_right integrable_distr_eq)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2472
  next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2473
    case (lim f s)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2474
    have [measurable]: "f \<in> borel_measurable N" "\<And>i. s i \<in> borel_measurable N"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2475
      using lim(1,5)[THEN borel_measurable_integrable] by auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2476
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2477
    show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2478
    proof (rule LIMSEQ_unique)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2479
      show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. f (g x))"
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2480
      proof (rule integral_dominated_convergence)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2481
        show "integrable M (\<lambda>x. 2 * norm (f (g x)))"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2482
          using lim by (auto simp: integrable_distr_eq)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2483
        show "AE x in M. (\<lambda>i. s i (g x)) \<longlonglongrightarrow> f (g x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2484
          using lim(3) g[THEN measurable_space] by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2485
        show "\<And>i. AE x in M. norm (s i (g x)) \<le> 2 * norm (f (g x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2486
          using lim(4) g[THEN measurable_space] by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2487
      qed auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2488
      show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) \<longlonglongrightarrow> integral\<^sup>L (distr M N g) f"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2489
        unfolding lim(2)[symmetric]
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
  2490
        by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  2491
           (insert lim(3-5), auto)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2492
    qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2493
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2494
qed (simp add: f g integrable_distr_eq)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2495
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2496
lemma has_bochner_integral_distr:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2497
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2498
  shows "f \<in> borel_measurable N \<Longrightarrow> g \<in> measurable M N \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2499
    has_bochner_integral M (\<lambda>x. f (g x)) x \<Longrightarrow> has_bochner_integral (distr M N g) f x"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2500
  by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr)
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68798
diff changeset
  2501
69683
8b3458ca0762 subsection is always %important
immler
parents: 69652
diff changeset
  2502
subsection \<open>Lebesgue integration on \<^const>\<open>count_space\<close>\<close>
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68798
diff changeset
  2503
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2504
lemma integrable_count_space:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2505
  fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2506
  shows "finite X \<Longrightarrow> integrable (count_space X) f"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2507
  by (auto simp: nn_integral_count_space integrable_iff_bounded)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2508
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2509
lemma measure_count_space[simp]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2510
  "B \<subseteq> A \<Longrightarrow> finite B \<Longrightarrow> measure (count_space A) B = card B"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2511
  unfolding measure_def by (subst emeasure_count_space ) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2512
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2513
lemma lebesgue_integral_count_space_finite_support:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2514
  assumes f: "finite {a\<in>A. f a \<noteq> 0}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2515
  shows "(\<integral>x. f x \<partial>count_space A) = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. f a)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2516
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2517
  have eq: "\<And>x. x \<in> A \<Longrightarrow> (\<Sum>a | x = a \<and> a \<in> A \<and> f a \<noteq> 0. f a) = (\<Sum>x\<in>{x}. f x)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2518
    by (intro sum.mono_neutral_cong_left) auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2519
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2520
  have "(\<integral>x. f x \<partial>count_space A) = (\<integral>x. (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. indicator {a} x *\<^sub>R f a) \<partial>count_space A)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2521
    by (intro integral_cong refl) (simp add: f eq)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2522
  also have "\<dots> = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. measure (count_space A) {a} *\<^sub>R f a)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2523
    by (subst integral_sum) (auto intro!: sum.cong)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2524
  finally show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2525
    by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2526
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2527
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2528
lemma lebesgue_integral_count_space_finite: "finite A \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2529
  by (subst lebesgue_integral_count_space_finite_support)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2530
     (auto intro!: sum.mono_neutral_cong_left)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2531
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2532
lemma integrable_count_space_nat_iff:
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2533
  fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2534
  shows "integrable (count_space UNIV) f \<longleftrightarrow> summable (\<lambda>x. norm (f x))"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2535
  by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2536
           intro:  summable_suminf_not_top)
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2537
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2538
lemma sums_integral_count_space_nat:
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2539
  fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2540
  assumes *: "integrable (count_space UNIV) f"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2541
  shows "f sums (integral\<^sup>L (count_space UNIV) f)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2542
proof -
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2543
  let ?f = "\<lambda>n i. indicator {n} i *\<^sub>R f i"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2544
  have f': "\<And>n i. ?f n i = indicator {n} i *\<^sub>R f n"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2545
    by (auto simp: fun_eq_iff split: split_indicator)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2546
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2547
  have "(\<lambda>i. \<integral>n. ?f i n \<partial>count_space UNIV) sums \<integral> n. (\<Sum>i. ?f i n) \<partial>count_space UNIV"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2548
  proof (rule sums_integral)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2549
    show "\<And>i. integrable (count_space UNIV) (?f i)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2550
      using * by (intro integrable_mult_indicator) auto
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2551
    show "AE n in count_space UNIV. summable (\<lambda>i. norm (?f i n))"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2552
      using summable_finite[of "{n}" "\<lambda>i. norm (?f i n)" for n] by simp
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2553
    show "summable (\<lambda>i. \<integral> n. norm (?f i n) \<partial>count_space UNIV)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2554
      using * by (subst f') (simp add: integrable_count_space_nat_iff)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2555
  qed
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2556
  also have "(\<integral> n. (\<Sum>i. ?f i n) \<partial>count_space UNIV) = (\<integral>n. f n \<partial>count_space UNIV)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2557
    using suminf_finite[of "{n}" "\<lambda>i. ?f i n" for n] by (auto intro!: integral_cong)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2558
  also have "(\<lambda>i. \<integral>n. ?f i n \<partial>count_space UNIV) = f"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2559
    by (subst f') simp
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2560
  finally show ?thesis .
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2561
qed
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2562
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2563
lemma integral_count_space_nat:
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2564
  fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2565
  shows "integrable (count_space UNIV) f \<Longrightarrow> integral\<^sup>L (count_space UNIV) f = (\<Sum>x. f x)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2566
  using sums_integral_count_space_nat by (rule sums_unique)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2567
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2568
lemma integrable_bij_count_space:
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2569
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2570
  assumes g: "bij_betw g A B"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2571
  shows "integrable (count_space A) (\<lambda>x. f (g x)) \<longleftrightarrow> integrable (count_space B) f"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2572
  unfolding integrable_iff_bounded by (subst nn_integral_bij_count_space[OF g]) auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2573
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2574
lemma integral_bij_count_space:
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2575
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2576
  assumes g: "bij_betw g A B"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2577
  shows "integral\<^sup>L (count_space A) (\<lambda>x. f (g x)) = integral\<^sup>L (count_space B) f"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2578
  using g[THEN bij_betw_imp_funcset]
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2579
  apply (subst distr_bij_count_space[OF g, symmetric])
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2580
  apply (intro integral_distr[symmetric])
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2581
  apply auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2582
  done
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2583
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2584
lemma has_bochner_integral_count_space_nat:
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2585
  fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2586
  shows "has_bochner_integral (count_space UNIV) f x \<Longrightarrow> f sums x"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2587
  unfolding has_bochner_integral_iff by (auto intro!: sums_integral_count_space_nat)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63941
diff changeset
  2588
69683
8b3458ca0762 subsection is always %important
immler
parents: 69652
diff changeset
  2589
subsection \<open>Point measure\<close>
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68798
diff changeset
  2590
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2591
lemma lebesgue_integral_point_measure_finite:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2592
  fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2593
  shows "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2594
    integral\<^sup>L (point_measure A f) g = (\<Sum>a\<in>A. f a *\<^sub>R g a)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2595
  by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2596
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2597
proposition integrable_point_measure_finite:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2598
  fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and f :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2599
  shows "finite A \<Longrightarrow> integrable (point_measure A f) g"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2600
  unfolding point_measure_def
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2601
  apply (subst density_cong[where f'="\<lambda>x. ennreal (max 0 (f x))"])
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2602
  apply (auto split: split_max simp: ennreal_neg)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2603
  apply (subst integrable_density)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2604
  apply (auto simp: AE_count_space integrable_count_space)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2605
  done
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2606
69683
8b3458ca0762 subsection is always %important
immler
parents: 69652
diff changeset
  2607
subsection \<open>Lebesgue integration on \<^const>\<open>null_measure\<close>\<close>
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68798
diff changeset
  2608
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2609
lemma has_bochner_integral_null_measure_iff[iff]:
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2610
  "has_bochner_integral (null_measure M) f 0 \<longleftrightarrow> f \<in> borel_measurable M"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2611
  by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def]
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2612
           intro!: exI[of _ "\<lambda>n x. 0"] simple_bochner_integrable.intros)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2613
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2614
lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \<longleftrightarrow> f \<in> borel_measurable M"
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2615
  by (auto simp add: integrable.simps)
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2616
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2617
lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0"
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2618
  by (cases "integrable (null_measure M) f")
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2619
     (auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  2620
69683
8b3458ca0762 subsection is always %important
immler
parents: 69652
diff changeset
  2621
subsection \<open>Legacy lemmas for the real-valued Lebesgue integral\<close>
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2622
theorem real_lebesgue_integral_def:
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  2623
  assumes f[measurable]: "integrable M f"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2624
  shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+x. f x \<partial>M) - enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2625
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2626
  have "integral\<^sup>L M f = integral\<^sup>L M (\<lambda>x. max 0 (f x) - max 0 (- f x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2627
    by (auto intro!: arg_cong[where f="integral\<^sup>L M"])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2628
  also have "\<dots> = integral\<^sup>L M (\<lambda>x. max 0 (f x)) - integral\<^sup>L M (\<lambda>x. max 0 (- f x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2629
    by (intro integral_diff integrable_max integrable_minus integrable_zero f)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2630
  also have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = enn2real (\<integral>\<^sup>+x. ennreal (f x) \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2631
    by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2632
  also have "integral\<^sup>L M (\<lambda>x. max 0 (- f x)) = enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2633
    by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2634
  finally show ?thesis .
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2635
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2636
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2637
theorem real_integrable_def:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2638
  "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2639
    (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2640
  unfolding integrable_iff_bounded
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2641
proof (safe del: notI)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2642
  assume *: "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2643
  have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2644
    by (intro nn_integral_mono) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2645
  also note *
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2646
  finally show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2647
    by simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2648
  have "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2649
    by (intro nn_integral_mono) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2650
  also note *
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2651
  finally show "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2652
    by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2653
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2654
  assume [measurable]: "f \<in> borel_measurable M"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2655
  assume fin: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2656
  have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (f x) + ennreal (- f x) \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2657
    by (intro nn_integral_cong) (auto simp: abs_real_def ennreal_neg)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2658
  also have"\<dots> = (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) + (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2659
    by (intro nn_integral_add) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2660
  also have "\<dots> < \<infinity>"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2661
    using fin by (auto simp: less_top)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2662
  finally show "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) < \<infinity>" .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2663
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2664
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2665
lemma integrableD[dest]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2666
  assumes "integrable M f"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2667
  shows "f \<in> borel_measurable M" "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2668
  using assms unfolding real_integrable_def by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2669
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2670
lemma integrableE:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2671
  assumes "integrable M f"
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63886
diff changeset
  2672
  obtains r q where "0 \<le> r" "0 \<le> q"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2673
    "(\<integral>\<^sup>+x. ennreal (f x)\<partial>M) = ennreal r"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2674
    "(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M) = ennreal q"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2675
    "f \<in> borel_measurable M" "integral\<^sup>L M f = r - q"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2676
  using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2677
  by (cases rule: ennreal2_cases[of "(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M)" "(\<integral>\<^sup>+x. ennreal (f x)\<partial>M)"]) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2678
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2679
lemma integral_monotone_convergence_nonneg:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2680
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2681
  assumes i: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2682
    and pos: "\<And>i. AE x in M. 0 \<le> f i x"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2683
    and lim: "AE x in M. (\<lambda>i. f i x) \<longlonglongrightarrow> u x"
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2684
    and ilim: "(\<lambda>i. integral\<^sup>L M (f i)) \<longlonglongrightarrow> x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2685
    and u: "u \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2686
  shows "integrable M u"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2687
  and "integral\<^sup>L M u = x"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2688
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2689
  have nn: "AE x in M. \<forall>i. 0 \<le> f i x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2690
    using pos unfolding AE_all_countable by auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2691
  with lim have u_nn: "AE x in M. 0 \<le> u x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2692
    by eventually_elim (auto intro: LIMSEQ_le_const)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2693
  have [simp]: "0 \<le> x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2694
    by (intro LIMSEQ_le_const[OF ilim] allI exI impI integral_nonneg_AE pos)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2695
  have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = (SUP n. (\<integral>\<^sup>+ x. ennreal (f n x) \<partial>M))"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2696
  proof (subst nn_integral_monotone_convergence_SUP_AE[symmetric])
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2697
    fix i
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2698
    from mono nn show "AE x in M. ennreal (f i x) \<le> ennreal (f (Suc i) x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2699
      by eventually_elim (auto simp: mono_def)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2700
    show "(\<lambda>x. ennreal (f i x)) \<in> borel_measurable M"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2701
      using i by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2702
  next
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2703
    show "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = \<integral>\<^sup>+ x. (SUP i. ennreal (f i x)) \<partial>M"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2704
      apply (rule nn_integral_cong_AE)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2705
      using lim mono nn u_nn
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2706
      apply eventually_elim
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2707
      apply (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2708
      done
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2709
  qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2710
  also have "\<dots> = ennreal x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2711
    using mono i nn unfolding nn_integral_eq_integral[OF i pos]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2712
    by (subst LIMSEQ_unique[OF LIMSEQ_SUP]) (auto simp: mono_def integral_nonneg_AE pos intro!: integral_mono_AE ilim)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2713
  finally have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = ennreal x" .
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2714
  moreover have "(\<integral>\<^sup>+ x. ennreal (- u x) \<partial>M) = 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2715
    using u u_nn by (subst nn_integral_0_iff_AE) (auto simp add: ennreal_neg)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2716
  ultimately show "integrable M u" "integral\<^sup>L M u = x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2717
    by (auto simp: real_integrable_def real_lebesgue_integral_def u)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2718
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2719
69739
nipkow
parents: 69722
diff changeset
  2720
lemma
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2721
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2722
  assumes f: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2723
  and lim: "AE x in M. (\<lambda>i. f i x) \<longlonglongrightarrow> u x"
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2724
  and ilim: "(\<lambda>i. integral\<^sup>L M (f i)) \<longlonglongrightarrow> x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2725
  and u: "u \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2726
  shows integrable_monotone_convergence: "integrable M u"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2727
    and integral_monotone_convergence: "integral\<^sup>L M u = x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2728
    and has_bochner_integral_monotone_convergence: "has_bochner_integral M u x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2729
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2730
  have 1: "\<And>i. integrable M (\<lambda>x. f i x - f 0 x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2731
    using f by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2732
  have 2: "AE x in M. mono (\<lambda>n. f n x - f 0 x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2733
    using mono by (auto simp: mono_def le_fun_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2734
  have 3: "\<And>n. AE x in M. 0 \<le> f n x - f 0 x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2735
    using mono by (auto simp: field_simps mono_def le_fun_def)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2736
  have 4: "AE x in M. (\<lambda>i. f i x - f 0 x) \<longlonglongrightarrow> u x - f 0 x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2737
    using lim by (auto intro!: tendsto_diff)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2738
  have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x \<partial>M)) \<longlonglongrightarrow> x - integral\<^sup>L M (f 0)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2739
    using f ilim by (auto intro!: tendsto_diff)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2740
  have 6: "(\<lambda>x. u x - f 0 x) \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2741
    using f[of 0] u by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2742
  note diff = integral_monotone_convergence_nonneg[OF 1 2 3 4 5 6]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2743
  have "integrable M (\<lambda>x. (u x - f 0 x) + f 0 x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2744
    using diff(1) f by (rule integrable_add)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2745
  with diff(2) f show "integrable M u" "integral\<^sup>L M u = x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2746
    by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2747
  then show "has_bochner_integral M u x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2748
    by (metis has_bochner_integral_integrable)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2749
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2750
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2751
lemma integral_norm_eq_0_iff:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2752
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2753
  assumes f[measurable]: "integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2754
  shows "(\<integral>x. norm (f x) \<partial>M) = 0 \<longleftrightarrow> emeasure M {x\<in>space M. f x \<noteq> 0} = 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2755
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2756
  have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>x. norm (f x) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2757
    using f by (intro nn_integral_eq_integral integrable_norm) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2758
  then have "(\<integral>x. norm (f x) \<partial>M) = 0 \<longleftrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) = 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2759
    by simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2760
  also have "\<dots> \<longleftrightarrow> emeasure M {x\<in>space M. ennreal (norm (f x)) \<noteq> 0} = 0"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2761
    by (intro nn_integral_0_iff) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2762
  finally show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2763
    by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2764
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2765
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2766
lemma integral_0_iff:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2767
  fixes f :: "'a \<Rightarrow> real"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2768
  shows "integrable M f \<Longrightarrow> (\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> emeasure M {x\<in>space M. f x \<noteq> 0} = 0"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2769
  using integral_norm_eq_0_iff[of M f] by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2770
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2771
lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\<lambda>x. a)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2772
  using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong add: less_top[symmetric])
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2773
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2774
lemma lebesgue_integral_const[simp]:
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2775
  fixes a :: "'a :: {banach, second_countable_topology}"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2776
  shows "(\<integral>x. a \<partial>M) = measure M (space M) *\<^sub>R a"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2777
proof -
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2778
  { assume "emeasure M (space M) = \<infinity>" "a \<noteq> 0"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2779
    then have ?thesis
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2780
      by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) }
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2781
  moreover
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2782
  { assume "a = 0" then have ?thesis by simp }
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2783
  moreover
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2784
  { assume "emeasure M (space M) \<noteq> \<infinity>"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2785
    interpret finite_measure M
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2786
      proof qed fact
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2787
    have "(\<integral>x. a \<partial>M) = (\<integral>x. indicator (space M) x *\<^sub>R a \<partial>M)"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2788
      by (intro integral_cong) auto
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2789
    also have "\<dots> = measure M (space M) *\<^sub>R a"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2790
      by (simp add: less_top[symmetric])
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2791
    finally have ?thesis . }
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2792
  ultimately show ?thesis by blast
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2793
qed
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57137
diff changeset
  2794
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2795
lemma (in finite_measure) integrable_const_bound:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2796
  fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2797
  shows "AE x in M. norm (f x) \<le> B \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2798
  apply (rule integrable_bound[OF integrable_const[of B], of f])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2799
  apply assumption
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2800
  apply (cases "0 \<le> B")
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2801
  apply auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2802
  done
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2803
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2804
lemma (in finite_measure) integral_bounded_eq_bound_then_AE:
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2805
  assumes "AE x in M. f x \<le> (c::real)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2806
    "integrable M f" "(\<integral>x. f x \<partial>M) = c * measure M (space M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2807
  shows "AE x in M. f x = c"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2808
  apply (rule integral_ineq_eq_0_then_AE) using assms by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64272
diff changeset
  2809
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2810
lemma integral_indicator_finite_real:
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2811
  fixes f :: "'a \<Rightarrow> real"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2812
  assumes [simp]: "finite A"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2813
  assumes [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2814
  assumes finite: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} < \<infinity>"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2815
  shows "(\<integral>x. f x * indicator A x \<partial>M) = (\<Sum>a\<in>A. f a * measure M {a})"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2816
proof -
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2817
  have "(\<integral>x. f x * indicator A x \<partial>M) = (\<integral>x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2818
  proof (intro integral_cong refl)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2819
    fix x show "f x * indicator A x = (\<Sum>a\<in>A. f a * indicator {a} x)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2820
      by (auto split: split_indicator simp: eq_commute[of x] cong: conj_cong)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2821
  qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2822
  also have "\<dots> = (\<Sum>a\<in>A. f a * measure M {a})"
71633
07bec530f02e cleaned proofs
nipkow
parents: 70532
diff changeset
  2823
    using finite by (subst integral_sum) (auto)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2824
  finally show ?thesis .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2825
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2826
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2827
lemma (in finite_measure) ennreal_integral_real:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2828
  assumes [measurable]: "f \<in> borel_measurable M"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2829
  assumes ae: "AE x in M. f x \<le> ennreal B" "0 \<le> B"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2830
  shows "ennreal (\<integral>x. enn2real (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2831
proof (subst nn_integral_eq_integral[symmetric])
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2832
  show "integrable M (\<lambda>x. enn2real (f x))"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
  2833
    using ae by (intro integrable_const_bound[where B=B]) (auto simp: enn2real_leI)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2834
  show "(\<integral>\<^sup>+ x. ennreal (enn2real (f x)) \<partial>M) = integral\<^sup>N M f"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2835
    using ae by (intro nn_integral_cong_AE) (auto simp: le_less_trans[OF _ ennreal_less_top])
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
  2836
qed auto
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2837
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2838
lemma (in finite_measure) integral_less_AE:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2839
  fixes X Y :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2840
  assumes int: "integrable M X" "integrable M Y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2841
  assumes A: "(emeasure M) A \<noteq> 0" "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> X x \<noteq> Y x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2842
  assumes gt: "AE x in M. X x \<le> Y x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2843
  shows "integral\<^sup>L M X < integral\<^sup>L M Y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2844
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2845
  have "integral\<^sup>L M X \<le> integral\<^sup>L M Y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2846
    using gt int by (intro integral_mono_AE) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2847
  moreover
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2848
  have "integral\<^sup>L M X \<noteq> integral\<^sup>L M Y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2849
  proof
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2850
    assume eq: "integral\<^sup>L M X = integral\<^sup>L M Y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2851
    have "integral\<^sup>L M (\<lambda>x. \<bar>Y x - X x\<bar>) = integral\<^sup>L M (\<lambda>x. Y x - X x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2852
      using gt int by (intro integral_cong_AE) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2853
    also have "\<dots> = 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2854
      using eq int by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2855
    finally have "(emeasure M) {x \<in> space M. Y x - X x \<noteq> 0} = 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2856
      using int by (simp add: integral_0_iff)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2857
    moreover
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2858
    have "(\<integral>\<^sup>+x. indicator A x \<partial>M) \<le> (\<integral>\<^sup>+x. indicator {x \<in> space M. Y x - X x \<noteq> 0} x \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  2859
      using A by (intro nn_integral_mono_AE) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2860
    then have "(emeasure M) A \<le> (emeasure M) {x \<in> space M. Y x - X x \<noteq> 0}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2861
      using int A by (simp add: integrable_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2862
    ultimately have "emeasure M A = 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2863
      by simp
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
  2864
    with \<open>(emeasure M) A \<noteq> 0\<close> show False by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2865
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2866
  ultimately show ?thesis by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2867
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2868
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2869
lemma (in finite_measure) integral_less_AE_space:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2870
  fixes X Y :: "'a \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2871
  assumes int: "integrable M X" "integrable M Y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2872
  assumes gt: "AE x in M. X x < Y x" "emeasure M (space M) \<noteq> 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2873
  shows "integral\<^sup>L M X < integral\<^sup>L M Y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2874
  using gt by (intro integral_less_AE[OF int, where A="space M"]) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2875
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2876
lemma tendsto_integral_at_top:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2877
  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59000
diff changeset
  2878
  assumes [measurable_cong]: "sets M = sets borel" and f[measurable]: "integrable M f"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2879
  shows "((\<lambda>y. \<integral> x. indicator {.. y} x *\<^sub>R f x \<partial>M) \<longlongrightarrow> \<integral> x. f x \<partial>M) at_top"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2880
proof (rule tendsto_at_topI_sequentially)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2881
  fix X :: "nat \<Rightarrow> real" assume "filterlim X at_top sequentially"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2882
  show "(\<lambda>n. \<integral>x. indicator {..X n} x *\<^sub>R f x \<partial>M) \<longlonglongrightarrow> integral\<^sup>L M f"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2883
  proof (rule integral_dominated_convergence)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2884
    show "integrable M (\<lambda>x. norm (f x))"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2885
      by (rule integrable_norm) fact
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2886
    show "AE x in M. (\<lambda>n. indicator {..X n} x *\<^sub>R f x) \<longlonglongrightarrow> f x"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2887
    proof
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2888
      fix x
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2889
      from \<open>filterlim X at_top sequentially\<close>
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2890
      have "eventually (\<lambda>n. x \<le> X n) sequentially"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2891
        unfolding filterlim_at_top_ge[where c=x] by auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2892
      then show "(\<lambda>n. indicator {..X n} x *\<^sub>R f x) \<longlonglongrightarrow> f x"
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  2893
        by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono)
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2894
    qed
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2895
    fix n show "AE x in M. norm (indicator {..X n} x *\<^sub>R f x) \<le> norm (f x)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2896
      by (auto split: split_indicator)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  2897
  qed auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2898
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2899
69739
nipkow
parents: 69722
diff changeset
  2900
lemma
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2901
  fixes f :: "real \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2902
  assumes M: "sets M = sets borel"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2903
  assumes nonneg: "AE x in M. 0 \<le> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2904
  assumes borel: "f \<in> borel_measurable borel"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2905
  assumes int: "\<And>y. integrable M (\<lambda>x. f x * indicator {.. y} x)"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2906
  assumes conv: "((\<lambda>y. \<integral> x. f x * indicator {.. y} x \<partial>M) \<longlongrightarrow> x) at_top"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2907
  shows has_bochner_integral_monotone_convergence_at_top: "has_bochner_integral M f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2908
    and integrable_monotone_convergence_at_top: "integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2909
    and integral_monotone_convergence_at_top:"integral\<^sup>L M f = x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2910
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2911
  from nonneg have "AE x in M. mono (\<lambda>n::nat. f x * indicator {..real n} x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2912
    by (auto split: split_indicator intro!: monoI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2913
  { fix x have "eventually (\<lambda>n. f x * indicator {..real n} x = f x) sequentially"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61897
diff changeset
  2914
      by (rule eventually_sequentiallyI[of "nat \<lceil>x\<rceil>"])
59587
8ea7b22525cb Removed the obsolete functions "natfloor" and "natceiling"
nipkow
parents: 59452
diff changeset
  2915
         (auto split: split_indicator simp: nat_le_iff ceiling_le_iff) }
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2916
  from filterlim_cong[OF refl refl this]
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2917
  have "AE x in M. (\<lambda>i. f x * indicator {..real i} x) \<longlonglongrightarrow> f x"
58729
e8ecc79aee43 add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents: 57514
diff changeset
  2918
    by simp
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2919
  have "(\<lambda>i. \<integral> x. f x * indicator {..real i} x \<partial>M) \<longlonglongrightarrow> x"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2920
    using conv filterlim_real_sequentially by (rule filterlim_compose)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2921
  have M_measure[simp]: "borel_measurable M = borel_measurable borel"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2922
    using M by (simp add: sets_eq_imp_space_eq measurable_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2923
  have "f \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2924
    using borel by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2925
  show "has_bochner_integral M f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2926
    by (rule has_bochner_integral_monotone_convergence) fact+
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2927
  then show "integrable M f" "integral\<^sup>L M f = x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2928
    by (auto simp: _has_bochner_integral_iff)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2929
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2930
69683
8b3458ca0762 subsection is always %important
immler
parents: 69652
diff changeset
  2931
subsection \<open>Product measure\<close>
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68798
diff changeset
  2932
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2933
lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2934
  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
  2935
  assumes [measurable]: "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2936
  shows "Measurable.pred N (\<lambda>x. integrable M (f x))"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2937
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2938
  have [simp]: "\<And>x. x \<in> space N \<Longrightarrow> integrable M (f x) \<longleftrightarrow> (\<integral>\<^sup>+y. norm (f x y) \<partial>M) < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2939
    unfolding integrable_iff_bounded by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2940
  show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2941
    by (simp cong: measurable_cong)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2942
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2943
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2944
lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2945
  "(\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M) \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2946
    {x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M) \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2947
    (\<lambda>x. measure M (A x)) \<in> borel_measurable N"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2948
  unfolding measure_def by (intro measurable_emeasure borel_measurable_enn2real) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2949
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2950
proposition (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2951
  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
  2952
  assumes f[measurable]: "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2953
  shows "(\<lambda>x. \<integral>y. f x y \<partial>M) \<in> borel_measurable N"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  2954
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2955
  from borel_measurable_implies_sequence_metric[OF f, of 0] guess s ..
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2956
  then have s: "\<And>i. simple_function (N \<Otimes>\<^sub>M M) (s i)"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2957
    "\<And>x y. x \<in> space N \<Longrightarrow> y \<in> space M \<Longrightarrow> (\<lambda>i. s i (x, y)) \<longlonglongrightarrow> f x y"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2958
    "\<And>i x y. x \<in> space N \<Longrightarrow> y \<in> space M \<Longrightarrow> norm (s i (x, y)) \<le> 2 * norm (f x y)"
62379
340738057c8c An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents: 62093
diff changeset
  2959
    by (auto simp: space_pair_measure)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2960
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2961
  have [measurable]: "\<And>i. s i \<in> borel_measurable (N \<Otimes>\<^sub>M M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2962
    by (rule borel_measurable_simple_function) fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2963
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2964
  have "\<And>i. s i \<in> measurable (N \<Otimes>\<^sub>M M) (count_space UNIV)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2965
    by (rule measurable_simple_function) fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2966
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
  2967
  define f' where [abs_def]: "f' i x =
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
  2968
    (if integrable M (f x) then simple_bochner_integral M (\<lambda>y. s i (x, y)) else 0)" for i x
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2969
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2970
  { fix i x assume "x \<in> space N"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2971
    then have "simple_bochner_integral M (\<lambda>y. s i (x, y)) =
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2972
      (\<Sum>z\<in>s i ` (space N \<times> space M). measure M {y \<in> space M. s i (x, y) = z} *\<^sub>R z)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2973
      using s(1)[THEN simple_functionD(1)]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2974
      unfolding simple_bochner_integral_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2975
      by (intro sum.mono_neutral_cong_left)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2976
         (auto simp: eq_commute space_pair_measure image_iff cong: conj_cong) }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2977
  note eq = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2978
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2979
  show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2980
  proof (rule borel_measurable_LIMSEQ_metric)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2981
    fix i show "f' i \<in> borel_measurable N"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2982
      unfolding f'_def by (simp_all add: eq cong: measurable_cong if_cong)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2983
  next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2984
    fix x assume x: "x \<in> space N"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2985
    { assume int_f: "integrable M (f x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2986
      have int_2f: "integrable M (\<lambda>y. 2 * norm (f x y))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2987
        by (intro integrable_norm integrable_mult_right int_f)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2988
      have "(\<lambda>i. integral\<^sup>L M (\<lambda>y. s i (x, y))) \<longlonglongrightarrow> integral\<^sup>L M (f x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2989
      proof (rule integral_dominated_convergence)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2990
        from int_f show "f x \<in> borel_measurable M" by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2991
        show "\<And>i. (\<lambda>y. s i (x, y)) \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2992
          using x by simp
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2993
        show "AE xa in M. (\<lambda>i. s i (x, xa)) \<longlonglongrightarrow> f x xa"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2994
          using x s(2) by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2995
        show "\<And>i. AE xa in M. norm (s i (x, xa)) \<le> 2 * norm (f x xa)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2996
          using x s(3) by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2997
      qed fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2998
      moreover
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  2999
      { fix i
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3000
        have "simple_bochner_integrable M (\<lambda>y. s i (x, y))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3001
        proof (rule simple_bochner_integrableI_bounded)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3002
          have "(\<lambda>y. s i (x, y)) ` space M \<subseteq> s i ` (space N \<times> space M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3003
            using x by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3004
          then show "simple_function M (\<lambda>y. s i (x, y))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3005
            using simple_functionD(1)[OF s(1), of i] x
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3006
            by (intro simple_function_borel_measurable)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3007
               (auto simp: space_pair_measure dest: finite_subset)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3008
          have "(\<integral>\<^sup>+ y. ennreal (norm (s i (x, y))) \<partial>M) \<le> (\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  3009
            using x s by (intro nn_integral_mono) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3010
          also have "(\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M) < \<infinity>"
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  3011
            using int_2f unfolding integrable_iff_bounded by simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3012
          finally show "(\<integral>\<^sup>+ xa. ennreal (norm (s i (x, xa))) \<partial>M) < \<infinity>" .
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3013
        qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3014
        then have "integral\<^sup>L M (\<lambda>y. s i (x, y)) = simple_bochner_integral M (\<lambda>y. s i (x, y))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3015
          by (rule simple_bochner_integrable_eq_integral[symmetric]) }
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  3016
      ultimately have "(\<lambda>i. simple_bochner_integral M (\<lambda>y. s i (x, y))) \<longlonglongrightarrow> integral\<^sup>L M (f x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3017
        by simp }
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3018
    then
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  3019
    show "(\<lambda>i. f' i x) \<longlonglongrightarrow> integral\<^sup>L M (f x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3020
      unfolding f'_def
58729
e8ecc79aee43 add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents: 57514
diff changeset
  3021
      by (cases "integrable M (f x)") (simp_all add: not_integrable_integral_eq)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3022
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3023
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3024
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3025
lemma (in pair_sigma_finite) integrable_product_swap:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3026
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3027
  assumes "integrable (M1 \<Otimes>\<^sub>M M2) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3028
  shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3029
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60585
diff changeset
  3030
  interpret Q: pair_sigma_finite M2 M1 ..
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3031
  have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3032
  show ?thesis unfolding *
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3033
    by (rule integrable_distr[OF measurable_pair_swap'])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3034
       (simp add: distr_pair_swap[symmetric] assms)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3035
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3036
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3037
lemma (in pair_sigma_finite) integrable_product_swap_iff:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3038
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3039
  shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable (M1 \<Otimes>\<^sub>M M2) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3040
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60585
diff changeset
  3041
  interpret Q: pair_sigma_finite M2 M1 ..
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3042
  from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3043
  show ?thesis by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3044
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3045
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3046
lemma (in pair_sigma_finite) integral_product_swap:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3047
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3048
  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3049
  shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3050
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3051
  have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3052
  show ?thesis unfolding *
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3053
    by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3054
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3055
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3056
theorem (in pair_sigma_finite) Fubini_integrable:
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3057
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3058
  assumes f[measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3059
    and integ1: "integrable M1 (\<lambda>x. \<integral> y. norm (f (x, y)) \<partial>M2)"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3060
    and integ2: "AE x in M1. integrable M2 (\<lambda>y. f (x, y))"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3061
  shows "integrable (M1 \<Otimes>\<^sub>M M2) f"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3062
proof (rule integrableI_bounded)
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3063
  have "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. norm (f (x, y)) \<partial>M2) \<partial>M1)"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3064
    by (simp add: M2.nn_integral_fst [symmetric])
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3065
  also have "\<dots> = (\<integral>\<^sup>+ x. \<bar>\<integral>y. norm (f (x, y)) \<partial>M2\<bar> \<partial>M1)"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3066
    apply (intro nn_integral_cong_AE)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3067
    using integ2
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3068
  proof eventually_elim
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3069
    fix x assume "integrable M2 (\<lambda>y. f (x, y))"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3070
    then have f: "integrable M2 (\<lambda>y. norm (f (x, y)))"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3071
      by simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3072
    then have "(\<integral>\<^sup>+y. ennreal (norm (f (x, y))) \<partial>M2) = ennreal (LINT y|M2. norm (f (x, y)))"
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3073
      by (rule nn_integral_eq_integral) simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3074
    also have "\<dots> = ennreal \<bar>LINT y|M2. norm (f (x, y))\<bar>"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3075
      using f by simp
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3076
    finally show "(\<integral>\<^sup>+y. ennreal (norm (f (x, y))) \<partial>M2) = ennreal \<bar>LINT y|M2. norm (f (x, y))\<bar>" .
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3077
  qed
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3078
  also have "\<dots> < \<infinity>"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3079
    using integ1 by (simp add: integrable_iff_bounded integral_nonneg_AE)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3080
  finally show "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) < \<infinity>" .
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3081
qed fact
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
  3082
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3083
lemma (in pair_sigma_finite) emeasure_pair_measure_finite:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3084
  assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" and finite: "emeasure (M1 \<Otimes>\<^sub>M M2) A < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3085
  shows "AE x in M1. emeasure M2 {y\<in>space M2. (x, y) \<in> A} < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3086
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3087
  from M2.emeasure_pair_measure_alt[OF A] finite
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3088
  have "(\<integral>\<^sup>+ x. emeasure M2 (Pair x -` A) \<partial>M1) \<noteq> \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3089
    by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3090
  then have "AE x in M1. emeasure M2 (Pair x -` A) \<noteq> \<infinity>"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  3091
    by (rule nn_integral_PInf_AE[rotated]) (intro M2.measurable_emeasure_Pair A)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3092
  moreover have "\<And>x. x \<in> space M1 \<Longrightarrow> Pair x -` A = {y\<in>space M2. (x, y) \<in> A}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3093
    using sets.sets_into_space[OF A] by (auto simp: space_pair_measure)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3094
  ultimately show ?thesis by (auto simp: less_top)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3095
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3096
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3097
lemma (in pair_sigma_finite) AE_integrable_fst':
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3098
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3099
  assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3100
  shows "AE x in M1. integrable M2 (\<lambda>y. f (x, y))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3101
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3102
  have "(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1) = (\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2))"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  3103
    by (rule M2.nn_integral_fst) simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3104
  also have "(\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2)) \<noteq> \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3105
    using f unfolding integrable_iff_bounded by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3106
  finally have "AE x in M1. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  3107
    by (intro nn_integral_PInf_AE M2.borel_measurable_nn_integral )
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3108
       (auto simp: measurable_split_conv)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3109
  with AE_space show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3110
    by eventually_elim
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3111
       (auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]] less_top)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3112
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3113
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3114
lemma (in pair_sigma_finite) integrable_fst':
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3115
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3116
  assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3117
  shows "integrable M1 (\<lambda>x. \<integral>y. f (x, y) \<partial>M2)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3118
  unfolding integrable_iff_bounded
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3119
proof
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3120
  show "(\<lambda>x. \<integral> y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3121
    by (rule M2.borel_measurable_lebesgue_integral) simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3122
  have "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) \<le> (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3123
    using AE_integrable_fst'[OF f] by (auto intro!: nn_integral_mono_AE integral_norm_bound_ennreal)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3124
  also have "(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1) = (\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2))"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  3125
    by (rule M2.nn_integral_fst) simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3126
  also have "(\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2)) < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3127
    using f unfolding integrable_iff_bounded by simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3128
  finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) < \<infinity>" .
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3129
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3130
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3131
proposition (in pair_sigma_finite) integral_fst':
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3132
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3133
  assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3134
  shows "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3135
using f proof induct
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3136
  case (base A c)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3137
  have A[measurable]: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" by fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3138
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3139
  have eq: "\<And>x y. x \<in> space M1 \<Longrightarrow> indicator A (x, y) = indicator {y\<in>space M2. (x, y) \<in> A} y"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3140
    using sets.sets_into_space[OF A] by (auto split: split_indicator simp: space_pair_measure)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3141
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3142
  have int_A: "integrable (M1 \<Otimes>\<^sub>M M2) (indicator A :: _ \<Rightarrow> real)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3143
    using base by (rule integrable_real_indicator)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3144
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3145
  have "(\<integral> x. \<integral> y. indicator A (x, y) *\<^sub>R c \<partial>M2 \<partial>M1) = (\<integral>x. measure M2 {y\<in>space M2. (x, y) \<in> A} *\<^sub>R c \<partial>M1)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3146
  proof (intro integral_cong_AE, simp, simp)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3147
    from AE_integrable_fst'[OF int_A] AE_space
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3148
    show "AE x in M1. (\<integral>y. indicator A (x, y) *\<^sub>R c \<partial>M2) = measure M2 {y\<in>space M2. (x, y) \<in> A} *\<^sub>R c"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3149
      by eventually_elim (simp add: eq integrable_indicator_iff)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3150
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3151
  also have "\<dots> = measure (M1 \<Otimes>\<^sub>M M2) A *\<^sub>R c"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3152
  proof (subst integral_scaleR_left)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3153
    have "(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3154
      (\<integral>\<^sup>+x. emeasure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3155
      using emeasure_pair_measure_finite[OF base]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3156
      by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ennreal_measure)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3157
    also have "\<dots> = emeasure (M1 \<Otimes>\<^sub>M M2) A"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3158
      using sets.sets_into_space[OF A]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3159
      by (subst M2.emeasure_pair_measure_alt)
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  3160
         (auto intro!: nn_integral_cong arg_cong[where f="emeasure M2"] simp: space_pair_measure)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3161
    finally have *: "(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) = emeasure (M1 \<Otimes>\<^sub>M M2) A" .
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3162
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3163
    from base * show "integrable M1 (\<lambda>x. measure M2 {y \<in> space M2. (x, y) \<in> A})"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3164
      by (simp add: integrable_iff_bounded)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3165
    then have "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3166
      (\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3167
      by (rule nn_integral_eq_integral[symmetric]) simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3168
    also note *
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3169
    finally show "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) *\<^sub>R c = measure (M1 \<Otimes>\<^sub>M M2) A *\<^sub>R c"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3170
      using base by (simp add: emeasure_eq_ennreal_measure)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3171
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3172
  also have "\<dots> = (\<integral> a. indicator A a *\<^sub>R c \<partial>(M1 \<Otimes>\<^sub>M M2))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3173
    using base by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3174
  finally show ?case .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3175
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3176
  case (add f g)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3177
  then have [measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" "g \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3178
    by auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3179
  have "(\<integral> x. \<integral> y. f (x, y) + g (x, y) \<partial>M2 \<partial>M1) =
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3180
    (\<integral> x. (\<integral> y. f (x, y) \<partial>M2) + (\<integral> y. g (x, y) \<partial>M2) \<partial>M1)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3181
    apply (rule integral_cong_AE)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3182
    apply simp_all
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3183
    using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3184
    apply eventually_elim
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3185
    apply simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3186
    done
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3187
  also have "\<dots> = (\<integral> x. f x \<partial>(M1 \<Otimes>\<^sub>M M2)) + (\<integral> x. g x \<partial>(M1 \<Otimes>\<^sub>M M2))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3188
    using integrable_fst'[OF add(1)] integrable_fst'[OF add(3)] add(2,4) by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3189
  finally show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3190
    using add by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3191
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3192
  case (lim f s)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3193
  then have [measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" "\<And>i. s i \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3194
    by auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3195
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3196
  show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3197
  proof (rule LIMSEQ_unique)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  3198
    show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) \<longlonglongrightarrow> integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3199
    proof (rule integral_dominated_convergence)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3200
      show "integrable (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. 2 * norm (f x))"
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  3201
        using lim(5) by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3202
    qed (insert lim, auto)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  3203
    have "(\<lambda>i. \<integral> x. \<integral> y. s i (x, y) \<partial>M2 \<partial>M1) \<longlonglongrightarrow> \<integral> x. \<integral> y. f (x, y) \<partial>M2 \<partial>M1"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3204
    proof (rule integral_dominated_convergence)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3205
      have "AE x in M1. \<forall>i. integrable M2 (\<lambda>y. s i (x, y))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3206
        unfolding AE_all_countable using AE_integrable_fst'[OF lim(1)] ..
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3207
      with AE_space AE_integrable_fst'[OF lim(5)]
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  3208
      show "AE x in M1. (\<lambda>i. \<integral> y. s i (x, y) \<partial>M2) \<longlonglongrightarrow> \<integral> y. f (x, y) \<partial>M2"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3209
      proof eventually_elim
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3210
        fix x assume x: "x \<in> space M1" and
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3211
          s: "\<forall>i. integrable M2 (\<lambda>y. s i (x, y))" and f: "integrable M2 (\<lambda>y. f (x, y))"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  3212
        show "(\<lambda>i. \<integral> y. s i (x, y) \<partial>M2) \<longlonglongrightarrow> \<integral> y. f (x, y) \<partial>M2"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3213
        proof (rule integral_dominated_convergence)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3214
          show "integrable M2 (\<lambda>y. 2 * norm (f (x, y)))"
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 57025
diff changeset
  3215
             using f by auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  3216
          show "AE xa in M2. (\<lambda>i. s i (x, xa)) \<longlonglongrightarrow> f (x, xa)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3217
            using x lim(3) by (auto simp: space_pair_measure)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3218
          show "\<And>i. AE xa in M2. norm (s i (x, xa)) \<le> 2 * norm (f (x, xa))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3219
            using x lim(4) by (auto simp: space_pair_measure)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3220
        qed (insert x, measurable)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3221
      qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3222
      show "integrable M1 (\<lambda>x. (\<integral> y. 2 * norm (f (x, y)) \<partial>M2))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3223
        by (intro integrable_mult_right integrable_norm integrable_fst' lim)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3224
      fix i show "AE x in M1. norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral> y. 2 * norm (f (x, y)) \<partial>M2)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3225
        using AE_space AE_integrable_fst'[OF lim(1), of i] AE_integrable_fst'[OF lim(5)]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3226
      proof eventually_elim
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3227
        fix x assume x: "x \<in> space M1"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3228
          and s: "integrable M2 (\<lambda>y. s i (x, y))" and f: "integrable M2 (\<lambda>y. f (x, y))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3229
        from s have "norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral>\<^sup>+y. norm (s i (x, y)) \<partial>M2)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3230
          by (rule integral_norm_bound_ennreal)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3231
        also have "\<dots> \<le> (\<integral>\<^sup>+y. 2 * norm (f (x, y)) \<partial>M2)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  3232
          using x lim by (auto intro!: nn_integral_mono simp: space_pair_measure)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3233
        also have "\<dots> = (\<integral>y. 2 * norm (f (x, y)) \<partial>M2)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  3234
          using f by (intro nn_integral_eq_integral) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3235
        finally show "norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral> y. 2 * norm (f (x, y)) \<partial>M2)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3236
          by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3237
      qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3238
    qed simp_all
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  3239
    then show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) \<longlonglongrightarrow> \<integral> x. \<integral> y. f (x, y) \<partial>M2 \<partial>M1"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3240
      using lim by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3241
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3242
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3243
69739
nipkow
parents: 69722
diff changeset
  3244
lemma (in pair_sigma_finite)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3245
  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
  3246
  assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3247
  shows AE_integrable_fst: "AE x in M1. integrable M2 (\<lambda>y. f x y)" (is "?AE")
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3248
    and integrable_fst: "integrable M1 (\<lambda>x. \<integral>y. f x y \<partial>M2)" (is "?INT")
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3249
    and integral_fst: "(\<integral>x. (\<integral>y. f x y \<partial>M2) \<partial>M1) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). f x y)" (is "?EQ")
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3250
  using AE_integrable_fst'[OF f] integrable_fst'[OF f] integral_fst'[OF f] by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3251
69739
nipkow
parents: 69722
diff changeset
  3252
lemma (in pair_sigma_finite)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3253
  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
  3254
  assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3255
  shows AE_integrable_snd: "AE y in M2. integrable M1 (\<lambda>x. f x y)" (is "?AE")
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3256
    and integrable_snd: "integrable M2 (\<lambda>y. \<integral>x. f x y \<partial>M1)" (is "?INT")
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
  3257
    and integral_snd: "(\<integral>y. (\<integral>x. f x y \<partial>M1) \<partial>M2) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (case_prod f)" (is "?EQ")
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3258
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60585
diff changeset
  3259
  interpret Q: pair_sigma_finite M2 M1 ..
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3260
  have Q_int: "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x, y). f y x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3261
    using f unfolding integrable_product_swap_iff[symmetric] by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3262
  show ?AE  using Q.AE_integrable_fst'[OF Q_int] by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3263
  show ?INT using Q.integrable_fst'[OF Q_int] by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3264
  show ?EQ using Q.integral_fst'[OF Q_int]
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
  3265
    using integral_product_swap[of "case_prod f"] by simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3266
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3267
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3268
proposition (in pair_sigma_finite) Fubini_integral:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3269
  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: {banach, second_countable_topology}"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
  3270
  assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3271
  shows "(\<integral>y. (\<integral>x. f x y \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f x y \<partial>M2) \<partial>M1)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3272
  unfolding integral_snd[OF assms] integral_fst[OF assms] ..
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3273
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3274
lemma (in product_sigma_finite) product_integral_singleton:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3275
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3276
  shows "f \<in> borel_measurable (M i) \<Longrightarrow> (\<integral>x. f (x i) \<partial>Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3277
  apply (subst distr_singleton[symmetric])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3278
  apply (subst integral_distr)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3279
  apply simp_all
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3280
  done
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3281
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3282
lemma (in product_sigma_finite) product_integral_fold:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3283
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3284
  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3285
  and f: "integrable (Pi\<^sub>M (I \<union> J) M) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3286
  shows "integral\<^sup>L (Pi\<^sub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I J (x, y)) \<partial>Pi\<^sub>M J M) \<partial>Pi\<^sub>M I M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3287
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60585
diff changeset
  3288
  interpret I: finite_product_sigma_finite M I by standard fact
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60585
diff changeset
  3289
  interpret J: finite_product_sigma_finite M J by standard fact
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3290
  have "finite (I \<union> J)" using fin by auto
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60585
diff changeset
  3291
  interpret IJ: finite_product_sigma_finite M "I \<union> J" by standard fact
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60585
diff changeset
  3292
  interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" ..
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3293
  let ?M = "merge I J"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3294
  let ?f = "\<lambda>x. f (?M x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3295
  from f have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3296
    by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3297
  have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3298
    using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3299
  have f_int: "integrable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) ?f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3300
    by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3301
  show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3302
    apply (subst distr_merge[symmetric, OF IJ fin])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3303
    apply (subst integral_distr[OF measurable_merge f_borel])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3304
    apply (subst P.integral_fst'[symmetric, OF f_int])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3305
    apply simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3306
    done
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3307
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3308
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3309
lemma (in product_sigma_finite) product_integral_insert:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3310
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3311
  assumes I: "finite I" "i \<notin> I"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3312
    and f: "integrable (Pi\<^sub>M (insert i I) M) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3313
  shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^sub>M I M)"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3314
proof -
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3315
  have "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = integral\<^sup>L (Pi\<^sub>M (I \<union> {i}) M) f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3316
    by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3317
  also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) \<partial>Pi\<^sub>M I M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3318
    using f I by (intro product_integral_fold) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3319
  also have "\<dots> = (\<integral>x. (\<integral>y. f (x(i := y)) \<partial>M i) \<partial>Pi\<^sub>M I M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3320
  proof (rule integral_cong[OF refl], subst product_integral_singleton[symmetric])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3321
    fix x assume x: "x \<in> space (Pi\<^sub>M I M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3322
    have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3323
      using f by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3324
    show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61681
diff changeset
  3325
      using measurable_comp[OF measurable_component_update f_borel, OF x \<open>i \<notin> I\<close>]
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3326
      unfolding comp_def .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3327
    from x I show "(\<integral> y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>Pi\<^sub>M {i} M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3328
      by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def PiE_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3329
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3330
  finally show ?thesis .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3331
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3332
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3333
lemma (in product_sigma_finite) product_integrable_prod:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3334
  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> _::{real_normed_field,banach,second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3335
  assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3336
  shows "integrable (Pi\<^sub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f")
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3337
proof (unfold integrable_iff_bounded, intro conjI)
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60585
diff changeset
  3338
  interpret finite_product_sigma_finite M I by standard fact
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
  3339
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3340
  show "?f \<in> borel_measurable (Pi\<^sub>M I M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3341
    using assms by simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3342
  have "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3343
      (\<integral>\<^sup>+ x. (\<Prod>i\<in>I. ennreal (norm (f i (x i)))) \<partial>Pi\<^sub>M I M)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3344
    by (simp add: prod_norm prod_ennreal)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3345
  also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+ x. ennreal (norm (f i x)) \<partial>M i)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3346
    using assms by (intro product_nn_integral_prod) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3347
  also have "\<dots> < \<infinity>"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3348
    using integrable by (simp add: less_top[symmetric] ennreal_prod_eq_top integrable_iff_bounded)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3349
  finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) < \<infinity>" .
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3350
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3351
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3352
lemma (in product_sigma_finite) product_integral_prod:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3353
  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> _::{real_normed_field,banach,second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3354
  assumes "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3355
  shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>L (M i) (f i))"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3356
using assms proof induct
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3357
  case empty
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3358
  interpret finite_measure "Pi\<^sub>M {} M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3359
    by rule (simp add: space_PiM)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3360
  show ?case by (simp add: space_PiM measure_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3361
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3362
  case (insert i I)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3363
  then have iI: "finite (insert i I)" by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3364
  then have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3365
    integrable (Pi\<^sub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3366
    by (intro product_integrable_prod insert(4)) (auto intro: finite_subset)
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60585
diff changeset
  3367
  interpret I: finite_product_sigma_finite M I by standard fact
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3368
  have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3369
    using \<open>i \<notin> I\<close> by (auto intro!: prod.cong)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3370
  show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3371
    unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3372
    by (simp add: * insert prod subset_insertI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3373
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3374
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3375
lemma integrable_subalgebra:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3376
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3377
  assumes borel: "f \<in> borel_measurable N"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3378
  and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3379
  shows "integrable N f \<longleftrightarrow> integrable M f" (is ?P)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3380
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3381
  have "f \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3382
    using assms by (auto simp: measurable_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3383
  with assms show ?thesis
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  3384
    using assms by (auto simp: integrable_iff_bounded nn_integral_subalgebra)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3385
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3386
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3387
lemma integral_subalgebra:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3388
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3389
  assumes borel: "f \<in> borel_measurable N"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3390
  and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3391
  shows "integral\<^sup>L N f = integral\<^sup>L M f"
69652
3417a8f154eb updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69597
diff changeset
  3392
proof cases
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3393
  assume "integrable N f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3394
  then show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3395
  proof induct
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3396
    case base with assms show ?case by (auto simp: subset_eq measure_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3397
  next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3398
    case (add f g)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3399
    then have "(\<integral> a. f a + g a \<partial>N) = integral\<^sup>L M f + integral\<^sup>L M g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3400
      by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3401
    also have "\<dots> = (\<integral> a. f a + g a \<partial>M)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3402
      using add integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of g] by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3403
    finally show ?case .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3404
  next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3405
    case (lim f s)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3406
    then have M: "\<And>i. integrable M (s i)" "integrable M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3407
      using integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of "s i" for i] by simp_all
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3408
    show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3409
    proof (intro LIMSEQ_unique)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  3410
      show "(\<lambda>i. integral\<^sup>L N (s i)) \<longlonglongrightarrow> integral\<^sup>L N f"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3411
        apply (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3412
        using lim
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3413
        apply auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3414
        done
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  3415
      show "(\<lambda>i. integral\<^sup>L N (s i)) \<longlonglongrightarrow> integral\<^sup>L M f"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3416
        unfolding lim
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3417
        apply (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3418
        using lim M N(2)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3419
        apply auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3420
        done
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3421
    qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3422
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3423
qed (simp add: not_integrable_integral_eq integrable_subalgebra[OF assms])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3424
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3425
hide_const (open) simple_bochner_integral
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  3426
hide_const (open) simple_bochner_integrable
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3427
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff changeset
  3428
end