| author | wenzelm | 
| Sun, 13 Sep 2015 22:56:52 +0200 | |
| changeset 61169 | 4de9ff3ea29a | 
| parent 60585 | 48fdff264eb2 | 
| child 61424 | c3658c18b7bc | 
| permissions | -rw-r--r-- | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Probability/Bochner_Integration.thy  | 
| 
 
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  | 
|
| 58876 | 5  | 
section {* Bochner Integration for Vector-Valued Functions *}
 | 
| 
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  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
11  | 
text {*
 | 
| 
 
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  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
16  | 
*}  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
18  | 
lemma borel_measurable_implies_sequence_metric:  | 
| 
 
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"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
21  | 
shows "\<exists>F. (\<forall>i. simple_function M (F i)) \<and> (\<forall>x\<in>space M. (\<lambda>i. F i x) ----> f x) \<and>  | 
| 
 
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)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
23  | 
proof -  | 
| 
 
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  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
27  | 
def e \<equiv> "from_nat_into D"  | 
| 
 
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  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
31  | 
from `d \<in> D` D[of UNIV] `countable D` obtain i where "d = e i"  | 
| 
 
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  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
37  | 
  def A \<equiv> "\<lambda>m n. {x\<in>space M. dist (f x) (e n) < 1 / (Suc m) \<and> 1 / (Suc m) \<le> 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
 | 
38  | 
def B \<equiv> "\<lambda>m. disjointed (A m)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
40  | 
  def m \<equiv> "\<lambda>N x. Max {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
 | 
41  | 
def F \<equiv> "\<lambda>N::nat. \<lambda>x. 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)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
42  | 
then e (LEAST n. x \<in> B (m N x) n) else z"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
43  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
44  | 
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
 | 
45  | 
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
 | 
46  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
47  | 
  { fix m
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
48  | 
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
 | 
49  | 
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
 | 
50  | 
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
 | 
51  | 
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
 | 
52  | 
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
 | 
53  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
54  | 
  { 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
 | 
55  | 
    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
 | 
56  | 
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
 | 
57  | 
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
 | 
58  | 
by auto }  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
59  | 
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
 | 
60  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
61  | 
  { 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
 | 
62  | 
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
 | 
63  | 
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
 | 
64  | 
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
 | 
65  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
66  | 
show ?thesis  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
67  | 
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
 | 
68  | 
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
 | 
69  | 
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
 | 
70  | 
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
 | 
71  | 
    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
 | 
72  | 
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  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
74  | 
    { fix i
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
75  | 
      { 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
 | 
76  | 
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
 | 
77  | 
by (intro Least_le)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
78  | 
also assume "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
 | 
79  | 
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
 | 
80  | 
      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
 | 
81  | 
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
 | 
82  | 
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
 | 
83  | 
by (rule 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
 | 
84  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
85  | 
    { fix N i n x assume "i \<le> N" "n \<le> N" "x \<in> B i n"
 | 
| 60585 | 86  | 
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
 | 
87  | 
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
 | 
88  | 
moreover  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
89  | 
def L \<equiv> "LEAST 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
 | 
90  | 
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
 | 
91  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
92  | 
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
 | 
93  | 
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
 | 
94  | 
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
 | 
95  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
96  | 
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
 | 
97  | 
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
 | 
98  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
99  | 
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
 | 
100  | 
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
 | 
101  | 
note * = this  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
102  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
103  | 
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
 | 
104  | 
show "(\<lambda>i. F i 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
 | 
105  | 
proof cases  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
106  | 
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
 | 
107  | 
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
 | 
108  | 
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
 | 
109  | 
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
 | 
110  | 
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
 | 
111  | 
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
 | 
112  | 
using `f x = z` by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
113  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
114  | 
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
 | 
115  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
116  | 
show ?thesis  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
117  | 
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
 | 
118  | 
fix e :: real assume "0 < e"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
119  | 
with `f x \<noteq> z` obtain n where "1 / Suc n < e" "1 / Suc n < 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
 | 
120  | 
by (metis dist_nz order_less_trans neq_iff nat_approx_posE)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
121  | 
with `x\<in>space M` `f x \<noteq> z` have "x \<in> (\<Union>i. B n i)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
122  | 
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
 | 
123  | 
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
 | 
124  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
125  | 
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
 | 
126  | 
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
 | 
127  | 
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
 | 
128  | 
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
 | 
129  | 
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
 | 
130  | 
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
 | 
131  | 
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
 | 
132  | 
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
 | 
133  | 
by (auto 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
 | 
134  | 
also note `1 / Suc n < e`  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
135  | 
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
 | 
136  | 
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
 | 
137  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
138  | 
qed  | 
| 
 
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  | 
fix i  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
141  | 
    { 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
 | 
142  | 
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
 | 
143  | 
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
 | 
144  | 
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
 | 
145  | 
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
 | 
146  | 
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
 | 
147  | 
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
 | 
148  | 
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
 | 
149  | 
apply auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
150  | 
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
 | 
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  | 
done  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
153  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
154  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
155  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
156  | 
lemma real_indicator: "real (indicator A x :: ereal) = 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
 | 
157  | 
unfolding indicator_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
 | 
158  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
159  | 
lemma split_indicator_asm:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
160  | 
"P (indicator S x) \<longleftrightarrow> \<not> ((x \<in> S \<and> \<not> P 1) \<or> (x \<notin> S \<and> \<not> P 0))"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
161  | 
unfolding indicator_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
 | 
162  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
163  | 
lemma  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
164  | 
fixes f :: "'a \<Rightarrow> 'b::semiring_1" assumes "finite A"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
165  | 
  shows setsum_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)"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
166  | 
  and setsum_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)"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
167  | 
unfolding indicator_def  | 
| 57418 | 168  | 
using assms by (auto intro!: setsum.mono_neutral_cong_right split: split_if_asm)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
169  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
170  | 
lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
171  | 
  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
 | 
172  | 
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
 | 
173  | 
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
 | 
174  | 
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
 | 
175  | 
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)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
176  | 
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) ----> u x) \<Longrightarrow> P u"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
177  | 
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
 | 
178  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
179  | 
have "(\<lambda>x. ereal (u x)) \<in> borel_measurable M" using 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
 | 
180  | 
from borel_measurable_implies_simple_function_sequence'[OF this]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
181  | 
obtain U where U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i. \<infinity> \<notin> range (U i)" and  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
182  | 
sup: "\<And>x. (SUP i. U i x) = max 0 (ereal (u x))" and nn: "\<And>i x. 0 \<le> 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
 | 
183  | 
by blast  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
184  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
185  | 
def U' \<equiv> "\<lambda>i x. indicator (space M) x * real (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
 | 
186  | 
then have U'_sf[measurable]: "\<And>i. simple_function M (U' i)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
187  | 
using U by (auto intro!: simple_function_compose1[where g=real])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
188  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
189  | 
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
 | 
190  | 
proof (rule seq)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
191  | 
fix i show "U' i \<in> borel_measurable M" "\<And>x. 0 \<le> 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
 | 
192  | 
using U nn by (auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
193  | 
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
 | 
194  | 
intro!: borel_measurable_real_of_ereal real_of_ereal_pos borel_measurable_times  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
195  | 
simp: U'_def zero_le_mult_iff)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
196  | 
show "incseq U'"  | 
| 
 
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,3) nn  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
198  | 
by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
199  | 
intro!: real_of_ereal_positive_mono)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
200  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
201  | 
fix x assume x: "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
 | 
202  | 
have "(\<lambda>i. U i x) ----> (SUP i. 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
 | 
203  | 
using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_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
 | 
204  | 
moreover have "(\<lambda>i. U i x) = (\<lambda>i. ereal (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
 | 
205  | 
using x nn U(3) by (auto simp: fun_eq_iff U'_def ereal_real image_iff eq_commute)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
206  | 
moreover have "(SUP i. U i x) = ereal (u x)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
207  | 
using sup u(2) by (simp add: max_def)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
208  | 
ultimately show "(\<lambda>i. U' i x) ----> u x"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
209  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
210  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
211  | 
fix i  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
212  | 
have "U' i ` space M \<subseteq> real ` (U i ` space M)" "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
 | 
213  | 
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
 | 
214  | 
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
 | 
215  | 
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
 | 
216  | 
    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
 | 
217  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
218  | 
    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
 | 
219  | 
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
 | 
220  | 
have "\<And>x. x \<in> U' i ` space M \<Longrightarrow> 0 \<le> x"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
221  | 
using nn by (auto simp: U'_def real_of_ereal_pos)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
222  | 
    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
 | 
223  | 
proof induct  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
224  | 
      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
 | 
225  | 
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
 | 
226  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
227  | 
case (insert x F)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
228  | 
then show ?case  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
229  | 
by (auto intro!: add mult set setsum_nonneg split: split_indicator split_indicator_asm  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
230  | 
simp del: setsum_mult_indicator simp: setsum_nonneg_eq_0_iff )  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
231  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
232  | 
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
 | 
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  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
235  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
236  | 
lemma scaleR_cong_right:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
237  | 
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
 | 
238  | 
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
 | 
239  | 
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
 | 
240  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
241  | 
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
 | 
242  | 
  "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
 | 
243  | 
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
 | 
244  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
245  | 
lemma 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
 | 
246  | 
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
 | 
247  | 
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
 | 
248  | 
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
 | 
249  | 
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
 | 
250  | 
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
 | 
251  | 
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
 | 
252  | 
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
 | 
253  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
254  | 
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
 | 
255  | 
"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
 | 
256  | 
"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
 | 
257  | 
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
 | 
258  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
259  | 
  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>"
 | 
| 
 
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  | 
  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
 | 
262  | 
      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
 | 
263  | 
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
 | 
264  | 
  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
 | 
265  | 
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
 | 
266  | 
  finally show "emeasure M {y \<in> space M. p (f y) (g y) \<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
 | 
267  | 
using fin by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
268  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
269  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
270  | 
lemma simple_function_finite_support:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
271  | 
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
 | 
272  | 
  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
 | 
273  | 
proof cases  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
274  | 
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
 | 
275  | 
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
 | 
276  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
277  | 
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
 | 
278  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
279  | 
  def m \<equiv> "Min (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
 | 
280  | 
  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
 | 
281  | 
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
 | 
282  | 
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
 | 
283  | 
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
 | 
284  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
285  | 
  from m have "m * emeasure M {x\<in>space M. 0 \<noteq> f x} = 
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
286  | 
    (\<integral>\<^sup>+x. m * indicator {x\<in>space M. 0 \<noteq> f x} x \<partial>M)"
 | 
| 56996 | 287  | 
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
 | 
288  | 
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
 | 
289  | 
using AE_space  | 
| 56996 | 290  | 
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
 | 
291  | 
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
 | 
292  | 
    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
 | 
293  | 
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
 | 
294  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
295  | 
also note `\<dots> < \<infinity>`  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
296  | 
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
 | 
297  | 
using 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
 | 
298  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
299  | 
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
 | 
300  | 
  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
 | 
301  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
302  | 
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
 | 
303  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
304  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
305  | 
lemma 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
 | 
306  | 
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
 | 
307  | 
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
 | 
308  | 
proof  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
309  | 
  have "emeasure M {y \<in> space M. ereal (norm (f y)) \<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
 | 
310  | 
proof (rule simple_function_finite_support)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
311  | 
show "simple_function M (\<lambda>x. ereal (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
 | 
312  | 
using f by (rule simple_function_compose1)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
313  | 
show "(\<integral>\<^sup>+ y. ereal (norm (f y)) \<partial>M) < \<infinity>" by fact  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
314  | 
qed simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
315  | 
  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
 | 
316  | 
qed fact  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
317  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
318  | 
definition simple_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> 'b" where
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
319  | 
  "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
 | 
320  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
321  | 
lemma 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
 | 
322  | 
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
 | 
323  | 
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
 | 
324  | 
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
 | 
325  | 
  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
 | 
326  | 
(is "_ = ?r")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
327  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
328  | 
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
 | 
329  | 
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
 | 
330  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
331  | 
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
 | 
332  | 
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
 | 
333  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
334  | 
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
 | 
335  | 
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
 | 
336  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
337  | 
  { 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
 | 
338  | 
    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
 | 
339  | 
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
 | 
340  | 
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
 | 
341  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
342  | 
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
 | 
343  | 
(\<Sum>y\<in>f`space M. (\<Sum>z\<in>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
 | 
344  | 
      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
 | 
345  | 
unfolding simple_bochner_integral_def  | 
| 57418 | 346  | 
proof (safe intro!: setsum.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
 | 
347  | 
fix y assume y: "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
 | 
348  | 
    have [simp]: "g ` space M \<inter> {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
 | 
349  | 
        {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
 | 
350  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
351  | 
    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
 | 
352  | 
        (\<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
 | 
353  | 
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
 | 
354  | 
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
 | 
355  | 
    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
 | 
356  | 
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
 | 
357  | 
moreover  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
358  | 
    { 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
 | 
359  | 
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
 | 
360  | 
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
 | 
361  | 
      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
 | 
362  | 
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
 | 
363  | 
      then have "emeasure M {xa \<in> space M. g xa = g x} < \<infinity>"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
364  | 
using f by (auto simp: 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
 | 
365  | 
ultimately  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
366  | 
    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
 | 
367  | 
      (\<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)"
 | 
| 57418 | 368  | 
apply (simp add: setsum.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
 | 
369  | 
apply (subst measure_finite_Union[symmetric])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
370  | 
apply (auto simp: disjoint_family_on_def)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
371  | 
done  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
372  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
373  | 
also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>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
 | 
374  | 
      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))"
 | 
| 57418 | 375  | 
by (auto intro!: setsum.cong simp: scaleR_setsum_left)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
376  | 
also have "\<dots> = ?r"  | 
| 57418 | 377  | 
by (subst setsum.commute)  | 
378  | 
(auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_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
 | 
379  | 
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
 | 
380  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
381  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
382  | 
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
 | 
383  | 
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
 | 
384  | 
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
 | 
385  | 
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
 | 
386  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
387  | 
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
 | 
388  | 
    (\<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
 | 
389  | 
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
 | 
390  | 
(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
 | 
391  | 
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
 | 
392  | 
    (\<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
 | 
393  | 
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
 | 
394  | 
(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
 | 
395  | 
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
 | 
396  | 
    (\<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
 | 
397  | 
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
 | 
398  | 
(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
 | 
399  | 
ultimately show ?thesis  | 
| 57418 | 400  | 
by (simp add: setsum.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
 | 
401  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
402  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
403  | 
lemma (in linear) 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
 | 
404  | 
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
 | 
405  | 
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
 | 
406  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
407  | 
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
 | 
408  | 
    (\<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
 | 
409  | 
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
 | 
410  | 
(auto simp: simple_bochner_integrable_compose2[where p="\<lambda>x y. f x"] zero  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
411  | 
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
 | 
412  | 
also have "\<dots> = 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
 | 
413  | 
by (simp add: simple_bochner_integral_def setsum scaleR)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
414  | 
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
 | 
415  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
416  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
417  | 
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
 | 
418  | 
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
 | 
419  | 
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
 | 
420  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
421  | 
interpret linear uminus by unfold_locales auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
422  | 
from f show ?thesis  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
423  | 
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
 | 
424  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
425  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
426  | 
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
 | 
427  | 
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
 | 
428  | 
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
 | 
429  | 
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
 | 
430  | 
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
 | 
431  | 
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
 | 
432  | 
(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
 | 
433  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
434  | 
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
 | 
435  | 
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
 | 
436  | 
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
 | 
437  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
438  | 
have "norm (simple_bochner_integral M f) \<le>  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
439  | 
    (\<Sum>y\<in>f ` space M. norm (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
 | 
440  | 
unfolding simple_bochner_integral_def by (rule norm_setsum)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
441  | 
  also have "\<dots> = (\<Sum>y\<in>f ` space M. measure M {x \<in> space M. f x = y} *\<^sub>R norm y)"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
442  | 
by (simp add: measure_nonneg)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
443  | 
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
 | 
444  | 
using f  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
445  | 
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
 | 
446  | 
(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
 | 
447  | 
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
 | 
448  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
449  | 
|
| 56996 | 450  | 
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
 | 
451  | 
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
 | 
452  | 
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
 | 
453  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
454  | 
  { fix x y z have "(x \<noteq> 0 \<Longrightarrow> y = z) \<Longrightarrow> ereal x * y = ereal x * z"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
455  | 
by (cases "x = 0") (auto simp: zero_ereal_def[symmetric]) }  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
456  | 
note ereal_cong_mult = this  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
457  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
458  | 
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
 | 
459  | 
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
 | 
460  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
461  | 
  { fix y assume y: "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
 | 
462  | 
    have "ereal (measure M {x \<in> space M. f x = f y}) = emeasure 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
 | 
463  | 
proof (rule emeasure_eq_ereal_measure[symmetric])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
464  | 
      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
 | 
465  | 
using y 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
 | 
466  | 
      with f show "emeasure M {x \<in> space M. f x = f y} \<noteq> \<infinity>"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
467  | 
by (auto simp: 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
 | 
468  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
469  | 
    moreover have "{x \<in> space M. f x = f y} = (\<lambda>x. ereal (f x)) -` {ereal (f y)} \<inter> space M"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
470  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
471  | 
    ultimately have "ereal (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
 | 
472  | 
          emeasure M ((\<lambda>x. ereal (f x)) -` {ereal (f y)} \<inter> 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
 | 
473  | 
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
 | 
474  | 
unfolding simple_integral_def  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
475  | 
by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ereal (f x)" and v=real])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
476  | 
(auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases  | 
| 57418 | 477  | 
intro!: setsum.cong ereal_cong_mult  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
478  | 
simp: setsum_ereal[symmetric] times_ereal.simps(1)[symmetric] ac_simps  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
479  | 
simp del: setsum_ereal times_ereal.simps(1))  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
480  | 
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
 | 
481  | 
using f  | 
| 56996 | 482  | 
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
 | 
483  | 
(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
 | 
484  | 
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
 | 
485  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
486  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
487  | 
lemma simple_bochner_integral_bounded:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
488  | 
  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
 | 
489  | 
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
 | 
490  | 
assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
491  | 
shows "ereal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \<le>  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
492  | 
(\<integral>\<^sup>+ x. norm (f x - s x) \<partial>M) + (\<integral>\<^sup>+ x. norm (f 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
 | 
493  | 
(is "ereal (norm (?s - ?t)) \<le> ?S + ?T")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
494  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
495  | 
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
 | 
496  | 
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
 | 
497  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
498  | 
have "ereal (norm (?s - ?t)) = norm (simple_bochner_integral M (\<lambda>x. s 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
 | 
499  | 
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
 | 
500  | 
also have "\<dots> \<le> simple_bochner_integral M (\<lambda>x. norm (s 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
 | 
501  | 
using simple_bochner_integrable_compose2[of "op -" M "s" "t"] s t  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
502  | 
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
 | 
503  | 
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
 | 
504  | 
using simple_bochner_integrable_compose2[of "\<lambda>x y. norm (x - y)" M "s" "t"] s t  | 
| 56996 | 505  | 
by (auto 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
 | 
506  | 
also have "\<dots> \<le> (\<integral>\<^sup>+x. ereal (norm (f x - s x)) + ereal (norm (f x - t x)) \<partial>M)"  | 
| 56996 | 507  | 
by (auto intro!: nn_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
 | 
508  | 
(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
 | 
509  | 
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
 | 
510  | 
also have "\<dots> = ?S + ?T"  | 
| 56996 | 511  | 
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
 | 
512  | 
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
 | 
513  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
514  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
515  | 
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
 | 
516  | 
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
 | 
517  | 
"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
 | 
518  | 
(\<And>i. simple_bochner_integrable M (s i)) \<Longrightarrow>  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
519  | 
(\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) ----> 0 \<Longrightarrow>  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
520  | 
(\<lambda>i. simple_bochner_integral M (s i)) ----> x \<Longrightarrow>  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
521  | 
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
 | 
522  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
523  | 
lemma has_bochner_integral_cong:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
524  | 
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
 | 
525  | 
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
 | 
526  | 
unfolding has_bochner_integral.simps assms(1,3)  | 
| 56996 | 527  | 
using assms(2) by (simp cong: measurable_cong_strong nn_integral_cong_strong)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
528  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
529  | 
lemma has_bochner_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
 | 
530  | 
"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
 | 
531  | 
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
 | 
532  | 
unfolding has_bochner_integral.simps  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
533  | 
by (intro arg_cong[where f=Ex] ext conj_cong rev_conj_cong refl arg_cong[where f="\<lambda>x. x ----> 0"]  | 
| 56996 | 534  | 
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
 | 
535  | 
auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
536  | 
|
| 
59353
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
537  | 
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
 | 
538  | 
"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
 | 
539  | 
by (rule has_bochner_integral.cases)  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
540  | 
|
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
541  | 
lemma borel_measurable_has_bochner_integral'[measurable_dest]:  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
542  | 
"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
 | 
543  | 
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
 | 
544  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
545  | 
lemma 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
 | 
546  | 
"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
 | 
547  | 
by (rule has_bochner_integral.intros[where s="\<lambda>_. f"])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
548  | 
(auto 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
 | 
549  | 
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
 | 
550  | 
simp: zero_ereal_def[symmetric])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
551  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
552  | 
lemma 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
 | 
553  | 
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
 | 
554  | 
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
 | 
555  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
556  | 
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
 | 
557  | 
proof  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
558  | 
    have "{y \<in> space M. (indicator A y::real) \<noteq> 0} = A"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
559  | 
using sets.sets_into_space[OF `A\<in>sets M`] by (auto 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
 | 
560  | 
    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
 | 
561  | 
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
 | 
562  | 
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
 | 
563  | 
moreover have "simple_bochner_integral M (indicator A) = measure M A"  | 
| 56996 | 564  | 
using simple_bochner_integral_eq_nn_integral[OF sbi] A  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
565  | 
by (simp add: ereal_indicator emeasure_eq_ereal_measure)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
566  | 
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
 | 
567  | 
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
 | 
568  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
569  | 
|
| 57036 | 570  | 
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
 | 
571  | 
"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
 | 
572  | 
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
 | 
573  | 
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
 | 
574  | 
fix sf sg  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
575  | 
assume f_sf: "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - sf i 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
 | 
576  | 
assume g_sg: "(\<lambda>i. \<integral>\<^sup>+ x. norm (g x - sg i 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
 | 
577  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
578  | 
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
 | 
579  | 
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
 | 
580  | 
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
 | 
581  | 
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
 | 
582  | 
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
 | 
583  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
584  | 
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
 | 
585  | 
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
 | 
586  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
587  | 
show "(\<lambda>i. \<integral>\<^sup>+ x. (norm (f x + g x - (sf i x + sg i 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
 | 
588  | 
(is "?f ----> 0")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
589  | 
proof (rule tendsto_sandwich)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
590  | 
show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) ----> 0"  | 
| 56996 | 591  | 
by (auto simp: nn_integral_nonneg)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
592  | 
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
 | 
593  | 
(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
 | 
594  | 
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
 | 
595  | 
fix i have "?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) + ereal (norm (g x - sg i x)) \<partial>M)"  | 
| 56996 | 596  | 
by (auto intro!: nn_integral_mono norm_diff_triangle_ineq)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
597  | 
also have "\<dots> = ?g i"  | 
| 56996 | 598  | 
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
 | 
599  | 
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
 | 
600  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
601  | 
show "?g ----> 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
602  | 
using tendsto_add_ereal[OF _ _ f_sf g_sg] by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
603  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
604  | 
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
 | 
605  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
606  | 
lemma 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
 | 
607  | 
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
 | 
608  | 
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
 | 
609  | 
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
 | 
610  | 
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
 | 
611  | 
have [measurable]: "T \<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
 | 
612  | 
by (intro borel_measurable_continuous_on1 T.continuous_on continuous_on_id)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
613  | 
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
 | 
614  | 
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
 | 
615  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
616  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
617  | 
fix s assume f_s: "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - s i 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
 | 
618  | 
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
 | 
619  | 
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
 | 
620  | 
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
 | 
621  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
622  | 
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
 | 
623  | 
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
 | 
624  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
625  | 
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
 | 
626  | 
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
 | 
627  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
628  | 
show "(\<lambda>i. \<integral>\<^sup>+ x. norm (T (f x) - T (s i 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
 | 
629  | 
(is "?f ----> 0")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
630  | 
proof (rule tendsto_sandwich)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
631  | 
show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) ----> 0"  | 
| 56996 | 632  | 
by (auto simp: nn_integral_nonneg)  | 
| 
56993
 
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  | 
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
 | 
635  | 
(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
 | 
636  | 
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
 | 
637  | 
fix i have "?f i \<le> (\<integral>\<^sup>+ x. ereal K * norm (f x - s i x) \<partial>M)"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
638  | 
using K by (intro nn_integral_mono) (auto simp: ac_simps)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
639  | 
also have "\<dots> = ?g i"  | 
| 56996 | 640  | 
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
 | 
641  | 
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
 | 
642  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
643  | 
show "?g ----> 0"  | 
| 
59452
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
644  | 
using tendsto_cmult_ereal[OF _ f_s, of "ereal K"] 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
 | 
645  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
646  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
647  | 
assume "(\<lambda>i. simple_bochner_integral M (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
 | 
648  | 
with s show "(\<lambda>i. simple_bochner_integral M (\<lambda>x. T (s i 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
 | 
649  | 
by (auto intro!: T.tendsto simp: T.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
 | 
650  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
651  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
652  | 
lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\<lambda>x. 0) 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
653  | 
by (auto intro!: has_bochner_integral.intros[where s="\<lambda>_ _. 0"]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
654  | 
simp: zero_ereal_def[symmetric] 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
 | 
655  | 
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
 | 
656  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
657  | 
lemma has_bochner_integral_scaleR_left[intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
658  | 
"(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
 | 
659  | 
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
 | 
660  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
661  | 
lemma has_bochner_integral_scaleR_right[intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
662  | 
"(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
 | 
663  | 
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
 | 
664  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
665  | 
lemma has_bochner_integral_mult_left[intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
666  | 
  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
 | 
667  | 
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
 | 
668  | 
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
 | 
669  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
670  | 
lemma has_bochner_integral_mult_right[intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
671  | 
  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
 | 
672  | 
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
 | 
673  | 
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
 | 
674  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
675  | 
lemmas has_bochner_integral_divide =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
676  | 
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
 | 
677  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
678  | 
lemma has_bochner_integral_divide_zero[intro]:  | 
| 
59867
 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 
haftmann 
parents: 
59587 
diff
changeset
 | 
679  | 
  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
 | 
680  | 
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
 | 
681  | 
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
 | 
682  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
683  | 
lemma has_bochner_integral_inner_left[intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
684  | 
"(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
 | 
685  | 
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
 | 
686  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
687  | 
lemma has_bochner_integral_inner_right[intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
688  | 
"(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
 | 
689  | 
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
 | 
690  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
691  | 
lemmas 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
 | 
692  | 
has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
693  | 
lemmas has_bochner_integral_Re =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
694  | 
has_bochner_integral_bounded_linear[OF bounded_linear_Re]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
695  | 
lemmas has_bochner_integral_Im =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
696  | 
has_bochner_integral_bounded_linear[OF bounded_linear_Im]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
697  | 
lemmas has_bochner_integral_cnj =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
698  | 
has_bochner_integral_bounded_linear[OF bounded_linear_cnj]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
699  | 
lemmas has_bochner_integral_of_real =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
700  | 
has_bochner_integral_bounded_linear[OF bounded_linear_of_real]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
701  | 
lemmas has_bochner_integral_fst =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
702  | 
has_bochner_integral_bounded_linear[OF bounded_linear_fst]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
703  | 
lemmas has_bochner_integral_snd =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
704  | 
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
 | 
705  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
706  | 
lemma has_bochner_integral_indicator:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
707  | 
"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
 | 
708  | 
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
 | 
709  | 
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
 | 
710  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
711  | 
lemma has_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
 | 
712  | 
"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
 | 
713  | 
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
 | 
714  | 
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
 | 
715  | 
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
 | 
716  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
717  | 
lemma has_bochner_integral_setsum:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
718  | 
"(\<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
 | 
719  | 
has_bochner_integral M (\<lambda>x. \<Sum>i\<in>I. f i x) (\<Sum>i\<in>I. x i)"  | 
| 57036 | 720  | 
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
 | 
721  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
722  | 
lemma has_bochner_integral_implies_finite_norm:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
723  | 
"has_bochner_integral M f x \<Longrightarrow> (\<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
 | 
724  | 
proof (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
 | 
725  | 
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
 | 
726  | 
assume [measurable]: "f \<in> borel_measurable M" and s: "\<And>i. simple_bochner_integrable M (s i)" and  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
727  | 
lim_0: "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (f x - s i 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
 | 
728  | 
from order_tendstoD[OF lim_0, of "\<infinity>"]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
729  | 
obtain i where f_s_fin: "(\<integral>\<^sup>+ x. ereal (norm (f x - 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
 | 
730  | 
by (metis (mono_tags, lifting) eventually_False_sequentially eventually_elim1  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
731  | 
less_ereal.simps(4) zero_ereal_def)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
732  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
733  | 
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
 | 
734  | 
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
 | 
735  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
736  | 
  def m \<equiv> "if space M = {} then 0 else Max ((\<lambda>x. norm (s i x))`space M)"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
737  | 
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
 | 
738  | 
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
 | 
739  | 
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
 | 
740  | 
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
 | 
741  | 
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
 | 
742  | 
by (auto simp: m_def image_comp comp_def Max_ge_iff)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
743  | 
  then have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ereal m * indicator {x\<in>space M. s i x \<noteq> 0} x \<partial>M)"
 | 
| 56996 | 744  | 
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
 | 
745  | 
also have "\<dots> < \<infinity>"  | 
| 56996 | 746  | 
using s by (subst nn_integral_cmult_indicator) (auto simp: `0 \<le> m` 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
 | 
747  | 
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
 | 
748  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
749  | 
have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (norm (f x - s i x)) + ereal (norm (s i x)) \<partial>M)"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57447 
diff
changeset
 | 
750  | 
by (auto intro!: nn_integral_mono) (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
 | 
751  | 
also have "\<dots> = (\<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) + (\<integral>\<^sup>+x. norm (s i x) \<partial>M)"  | 
| 56996 | 752  | 
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
 | 
753  | 
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
 | 
754  | 
using s_fin f_s_fin by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
755  | 
finally show "(\<integral>\<^sup>+ x. ereal (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
 | 
756  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
757  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
758  | 
lemma 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
 | 
759  | 
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
 | 
760  | 
shows "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
 | 
761  | 
using assms proof  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
762  | 
fix s assume  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
763  | 
x: "(\<lambda>i. simple_bochner_integral M (s i)) ----> x" (is "?s ----> x") and  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
764  | 
s[simp]: "\<And>i. simple_bochner_integrable M (s i)" and  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
765  | 
lim: "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (f x - s i x)) \<partial>M) ----> 0" and  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
766  | 
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
 | 
767  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
768  | 
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
 | 
769  | 
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
 | 
770  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
771  | 
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
 | 
772  | 
proof (rule LIMSEQ_le)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
773  | 
show "(\<lambda>i. ereal (norm (?s i))) ----> norm x"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
774  | 
using x by (intro tendsto_intros lim_ereal[THEN iffD2])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
775  | 
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
 | 
776  | 
(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
 | 
777  | 
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
 | 
778  | 
fix n  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
779  | 
have "ereal (norm (?s n)) \<le> simple_bochner_integral M (\<lambda>x. norm (s n x))"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
780  | 
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
 | 
781  | 
also have "\<dots> = (\<integral>\<^sup>+x. norm (s n x) \<partial>M)"  | 
| 56996 | 782  | 
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
 | 
783  | 
(auto intro: s 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
 | 
784  | 
also have "\<dots> \<le> (\<integral>\<^sup>+x. ereal (norm (f x - s n x)) + norm (f x) \<partial>M)"  | 
| 56996 | 785  | 
by (auto intro!: nn_integral_mono)  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57447 
diff
changeset
 | 
786  | 
(metis add.commute norm_minus_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
 | 
787  | 
also have "\<dots> = ?t n"  | 
| 56996 | 788  | 
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
 | 
789  | 
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
 | 
790  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
791  | 
have "?t ----> 0 + (\<integral>\<^sup>+ x. ereal (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
 | 
792  | 
using has_bochner_integral_implies_finite_norm[OF i]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
793  | 
by (intro tendsto_add_ereal tendsto_const lim) auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
794  | 
then show "?t ----> \<integral>\<^sup>+ x. ereal (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
 | 
795  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
796  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
797  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
798  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
799  | 
lemma 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
 | 
800  | 
"has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M f y \<Longrightarrow> x = y"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
801  | 
proof (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
 | 
802  | 
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
 | 
803  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
804  | 
fix s t  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
805  | 
assume "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M) ----> 0" (is "?S ----> 0")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
806  | 
assume "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - t i x) \<partial>M) ----> 0" (is "?T ----> 0")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
807  | 
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
 | 
808  | 
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
 | 
809  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
810  | 
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
 | 
811  | 
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
 | 
812  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
813  | 
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
 | 
814  | 
let ?t = "\<lambda>i. simple_bochner_integral 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
 | 
815  | 
assume "?s ----> x" "?t ----> y"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
816  | 
then have "(\<lambda>i. norm (?s i - ?t i)) ----> norm (x - y)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
817  | 
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
 | 
818  | 
moreover  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
819  | 
have "(\<lambda>i. ereal (norm (?s i - ?t i))) ----> ereal 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
820  | 
proof (rule tendsto_sandwich)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
821  | 
show "eventually (\<lambda>i. 0 \<le> ereal (norm (?s i - ?t i))) sequentially" "(\<lambda>_. 0) ----> ereal 0"  | 
| 56996 | 822  | 
by (auto simp: nn_integral_nonneg zero_ereal_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
 | 
823  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
824  | 
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
 | 
825  | 
by (intro always_eventually allI simple_bochner_integral_bounded s t f)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
826  | 
show "(\<lambda>i. ?S i + ?T i) ----> ereal 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
827  | 
using tendsto_add_ereal[OF _ _ `?S ----> 0` `?T ----> 0`]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
828  | 
by (simp add: zero_ereal_def[symmetric])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
829  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
830  | 
then have "(\<lambda>i. norm (?s i - ?t i)) ----> 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
831  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
832  | 
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
 | 
833  | 
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
 | 
834  | 
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
 | 
835  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
836  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
837  | 
lemma 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
 | 
838  | 
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
 | 
839  | 
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
 | 
840  | 
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
 | 
841  | 
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
 | 
842  | 
using f  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
843  | 
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
 | 
844  | 
fix s assume "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (f x - s i 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
 | 
845  | 
also have "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (f x - s i x)) \<partial>M) = (\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (g x - s i 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
 | 
846  | 
using ae  | 
| 56996 | 847  | 
by (intro ext nn_integral_cong_AE, eventually_elim) simp  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
848  | 
finally show "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (g x - s i 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
 | 
849  | 
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
 | 
850  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
851  | 
lemma has_bochner_integral_eq_AE:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
852  | 
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
 | 
853  | 
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
 | 
854  | 
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
 | 
855  | 
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
 | 
856  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
857  | 
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
 | 
858  | 
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
 | 
859  | 
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
 | 
860  | 
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
 | 
861  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
862  | 
|
| 57137 | 863  | 
lemma simple_bochner_integrable_restrict_space:  | 
864  | 
fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"  | 
|
865  | 
assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M"  | 
|
866  | 
shows "simple_bochner_integrable (restrict_space M \<Omega>) f \<longleftrightarrow>  | 
|
867  | 
simple_bochner_integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"  | 
|
868  | 
by (simp add: simple_bochner_integrable.simps space_restrict_space  | 
|
869  | 
simple_function_restrict_space[OF \<Omega>] emeasure_restrict_space[OF \<Omega>] Collect_restrict  | 
|
870  | 
indicator_eq_0_iff conj_ac)  | 
|
871  | 
||
872  | 
lemma simple_bochner_integral_restrict_space:  | 
|
873  | 
fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"  | 
|
874  | 
assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M"  | 
|
875  | 
assumes f: "simple_bochner_integrable (restrict_space M \<Omega>) f"  | 
|
876  | 
shows "simple_bochner_integral (restrict_space M \<Omega>) f =  | 
|
877  | 
simple_bochner_integral M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"  | 
|
878  | 
proof -  | 
|
879  | 
have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x)`space M)"  | 
|
880  | 
using f simple_bochner_integrable_restrict_space[OF \<Omega>, of f]  | 
|
881  | 
by (simp add: simple_bochner_integrable.simps simple_function_def)  | 
|
882  | 
then show ?thesis  | 
|
883  | 
by (auto simp: space_restrict_space measure_restrict_space[OF \<Omega>(1)] le_infI2  | 
|
884  | 
simple_bochner_integral_def Collect_restrict  | 
|
885  | 
split: split_indicator split_indicator_asm  | 
|
| 57418 | 886  | 
intro!: setsum.mono_neutral_cong_left arg_cong2[where f=measure])  | 
| 57137 | 887  | 
qed  | 
888  | 
||
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
889  | 
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
 | 
890  | 
"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
 | 
891  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
892  | 
definition 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
 | 
893  | 
"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
 | 
894  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
895  | 
syntax  | 
| 
59357
 
f366643536cd
allow line breaks in integral notation
 
Andreas Lochbihler 
parents: 
59353 
diff
changeset
 | 
896  | 
  "_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
 | 
897  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
898  | 
translations  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
899  | 
"\<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
 | 
900  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
901  | 
lemma has_bochner_integral_integral_eq: "has_bochner_integral M f x \<Longrightarrow> 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
 | 
902  | 
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
 | 
903  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
904  | 
lemma 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
 | 
905  | 
"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
 | 
906  | 
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
 | 
907  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
908  | 
lemma 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
 | 
909  | 
"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
 | 
910  | 
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
 | 
911  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
912  | 
lemma simple_bochner_integrable_eq_integral:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
913  | 
"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
 | 
914  | 
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
 | 
915  | 
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
 | 
916  | 
|
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
917  | 
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
 | 
918  | 
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
 | 
919  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
920  | 
lemma integral_eq_cases:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
921  | 
"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
 | 
922  | 
(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
 | 
923  | 
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
 | 
924  | 
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
 | 
925  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
926  | 
lemma borel_measurable_integrable[measurable_dest]: "integrable M f \<Longrightarrow> 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
 | 
927  | 
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
 | 
928  | 
|
| 
59353
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
929  | 
lemma borel_measurable_integrable'[measurable_dest]:  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
930  | 
"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
 | 
931  | 
using borel_measurable_integrable[measurable] by measurable  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
932  | 
|
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
933  | 
lemma integrable_cong:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
934  | 
"M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> integrable N g"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
935  | 
using assms by (simp cong: has_bochner_integral_cong add: integrable.simps)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
936  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
937  | 
lemma integrable_cong_AE:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
938  | 
"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
 | 
939  | 
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
 | 
940  | 
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
 | 
941  | 
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
 | 
942  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
943  | 
lemma integral_cong:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
944  | 
"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"  | 
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
945  | 
using assms by (simp cong: has_bochner_integral_cong cong del: if_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
 | 
946  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
947  | 
lemma 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
 | 
948  | 
"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
 | 
949  | 
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
 | 
950  | 
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
 | 
951  | 
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
 | 
952  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
953  | 
lemma integrable_add[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x + g x)"  | 
| 57036 | 954  | 
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
 | 
955  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
956  | 
lemma integrable_zero[simp, intro]: "integrable M (\<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
 | 
957  | 
by (metis has_bochner_integral_zero integrable.simps)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
958  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
959  | 
lemma integrable_setsum[simp, intro]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> integrable M (\<lambda>x. \<Sum>i\<in>I. 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
 | 
960  | 
by (metis has_bochner_integral_setsum 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  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
962  | 
lemma integrable_indicator[simp, intro]: "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
 | 
963  | 
integrable M (\<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
 | 
964  | 
by (metis has_bochner_integral_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
 | 
965  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
966  | 
lemma integrable_real_indicator[simp, intro]: "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
 | 
967  | 
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
 | 
968  | 
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
 | 
969  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
970  | 
lemma integrable_diff[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> 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
 | 
971  | 
by (auto simp: integrable.simps intro: has_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
 | 
972  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
973  | 
lemma integrable_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. T (f x))"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
974  | 
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
 | 
975  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
976  | 
lemma integrable_scaleR_left[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f 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
 | 
977  | 
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
 | 
978  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
979  | 
lemma integrable_scaleR_right[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c *\<^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
 | 
980  | 
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
 | 
981  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
982  | 
lemma integrable_mult_left[simp, intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
983  | 
  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
 | 
984  | 
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
 | 
985  | 
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
 | 
986  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
987  | 
lemma integrable_mult_right[simp, intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
988  | 
  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
 | 
989  | 
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
 | 
990  | 
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
 | 
991  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
992  | 
lemma integrable_divide_zero[simp, intro]:  | 
| 
59867
 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 
haftmann 
parents: 
59587 
diff
changeset
 | 
993  | 
  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
 | 
994  | 
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
 | 
995  | 
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
 | 
996  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
997  | 
lemma integrable_inner_left[simp, intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
998  | 
"(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
 | 
999  | 
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
 | 
1000  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1001  | 
lemma integrable_inner_right[simp, intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1002  | 
"(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
 | 
1003  | 
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
 | 
1004  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1005  | 
lemmas integrable_minus[simp, intro] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1006  | 
integrable_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1007  | 
lemmas integrable_divide[simp, intro] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1008  | 
integrable_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
 | 
1009  | 
lemmas integrable_Re[simp, intro] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1010  | 
integrable_bounded_linear[OF bounded_linear_Re]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1011  | 
lemmas integrable_Im[simp, intro] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1012  | 
integrable_bounded_linear[OF bounded_linear_Im]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1013  | 
lemmas integrable_cnj[simp, intro] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1014  | 
integrable_bounded_linear[OF bounded_linear_cnj]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1015  | 
lemmas integrable_of_real[simp, intro] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1016  | 
integrable_bounded_linear[OF bounded_linear_of_real]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1017  | 
lemmas integrable_fst[simp, intro] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1018  | 
integrable_bounded_linear[OF bounded_linear_fst]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1019  | 
lemmas integrable_snd[simp, intro] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1020  | 
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
 | 
1021  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1022  | 
lemma integral_zero[simp]: "integral\<^sup>L M (\<lambda>x. 0) = 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1023  | 
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
 | 
1024  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1025  | 
lemma integral_add[simp]: "integrable M f \<Longrightarrow> 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
 | 
1026  | 
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
 | 
1027  | 
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
 | 
1028  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1029  | 
lemma integral_diff[simp]: "integrable M f \<Longrightarrow> 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
 | 
1030  | 
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
 | 
1031  | 
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
 | 
1032  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1033  | 
lemma integral_setsum[simp]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow>  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1034  | 
integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1035  | 
by (intro has_bochner_integral_integral_eq has_bochner_integral_setsum 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
 | 
1036  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1037  | 
lemma integral_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow>  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1038  | 
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
 | 
1039  | 
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
 | 
1040  | 
|
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1041  | 
lemma integral_bounded_linear':  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1042  | 
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
 | 
1043  | 
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
 | 
1044  | 
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
 | 
1045  | 
proof cases  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1046  | 
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
 | 
1047  | 
by simp  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1048  | 
next  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1049  | 
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
 | 
1050  | 
show ?thesis  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1051  | 
proof cases  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1052  | 
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
 | 
1053  | 
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
 | 
1054  | 
next  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1055  | 
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
 | 
1056  | 
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
 | 
1057  | 
proof  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1058  | 
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
 | 
1059  | 
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
 | 
1060  | 
show False  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1061  | 
by auto  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1062  | 
qed  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1063  | 
ultimately show ?thesis  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1064  | 
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
 | 
1065  | 
qed  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1066  | 
qed  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1067  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1068  | 
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"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1069  | 
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
 | 
1070  | 
|
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1071  | 
lemma integral_scaleR_right[simp]: "(\<integral> x. c *\<^sub>R f x \<partial>M) = c *\<^sub>R 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
 | 
1072  | 
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
 | 
1073  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1074  | 
lemma integral_mult_left[simp]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1075  | 
  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
 | 
1076  | 
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
 | 
1077  | 
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
 | 
1078  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1079  | 
lemma integral_mult_right[simp]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1080  | 
  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
 | 
1081  | 
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
 | 
1082  | 
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
 | 
1083  | 
|
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1084  | 
lemma integral_mult_left_zero[simp]:  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1085  | 
  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
 | 
1086  | 
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
 | 
1087  | 
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
 | 
1088  | 
|
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1089  | 
lemma integral_mult_right_zero[simp]:  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1090  | 
  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
 | 
1091  | 
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
 | 
1092  | 
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
 | 
1093  | 
|
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1094  | 
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"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1095  | 
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
 | 
1096  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1097  | 
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"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1098  | 
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
 | 
1099  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1100  | 
lemma integral_divide_zero[simp]:  | 
| 
59867
 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 
haftmann 
parents: 
59587 
diff
changeset
 | 
1101  | 
  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
 | 
1102  | 
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
 | 
1103  | 
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
 | 
1104  | 
|
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1105  | 
lemma integral_minus[simp]: "integral\<^sup>L M (\<lambda>x. - f x) = - 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
 | 
1106  | 
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
 | 
1107  | 
|
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1108  | 
lemma integral_complex_of_real[simp]: "integral\<^sup>L M (\<lambda>x. complex_of_real (f x)) = of_real (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
 | 
1109  | 
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
 | 
1110  | 
|
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1111  | 
lemma integral_cnj[simp]: "integral\<^sup>L M (\<lambda>x. cnj (f x)) = cnj (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
 | 
1112  | 
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
 | 
1113  | 
|
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1114  | 
lemmas integral_divide[simp] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1115  | 
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
 | 
1116  | 
lemmas integral_Re[simp] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1117  | 
integral_bounded_linear[OF bounded_linear_Re]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1118  | 
lemmas integral_Im[simp] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1119  | 
integral_bounded_linear[OF bounded_linear_Im]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1120  | 
lemmas integral_of_real[simp] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1121  | 
integral_bounded_linear[OF bounded_linear_of_real]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1122  | 
lemmas integral_fst[simp] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1123  | 
integral_bounded_linear[OF bounded_linear_fst]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1124  | 
lemmas integral_snd[simp] =  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1125  | 
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
 | 
1126  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1127  | 
lemma integral_norm_bound_ereal:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1128  | 
"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
 | 
1129  | 
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
 | 
1130  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1131  | 
lemma integrableI_sequence:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1132  | 
  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
 | 
1133  | 
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
 | 
1134  | 
assumes 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
 | 
1135  | 
assumes lim: "(\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) ----> 0" (is "?S ----> 0")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1136  | 
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
 | 
1137  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1138  | 
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
 | 
1139  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1140  | 
have "\<exists>x. ?s ----> x"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1141  | 
unfolding convergent_eq_cauchy  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1142  | 
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
 | 
1143  | 
fix e :: real assume "0 < e"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1144  | 
then have "0 < ereal (e / 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
 | 
1145  | 
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
 | 
1146  | 
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
 | 
1147  | 
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
 | 
1148  | 
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
 | 
1149  | 
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
 | 
1150  | 
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
 | 
1151  | 
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
 | 
1152  | 
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
 | 
1153  | 
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
 | 
1154  | 
by (intro simple_bochner_integral_bounded s f)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1155  | 
also have "\<dots> < ereal (e / 2) + e / 2"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1156  | 
using ereal_add_strict_mono[OF less_imp_le[OF M[OF n]] _ `?S n \<noteq> \<infinity>` M[OF m]]  | 
| 56996 | 1157  | 
by (auto simp: nn_integral_nonneg)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1158  | 
also have "\<dots> = e" by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1159  | 
finally show "dist (?s n) (?s m) < e"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1160  | 
by (simp add: dist_norm)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1161  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1162  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1163  | 
then obtain x where "?s ----> x" ..  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1164  | 
show ?thesis  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1165  | 
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
 | 
1166  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1167  | 
|
| 56996 | 1168  | 
lemma 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
 | 
1169  | 
  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
 | 
1170  | 
assumes [measurable]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1171  | 
"\<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
 | 
1172  | 
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
 | 
1173  | 
and w: "(\<integral>\<^sup>+x. w 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
 | 
1174  | 
and u': "AE x in M. (\<lambda>i. u i x) ----> u' x"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1175  | 
shows "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i 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
 | 
1176  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1177  | 
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
 | 
1178  | 
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
 | 
1179  | 
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
 | 
1180  | 
proof (eventually_elim, intro allI)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1181  | 
fix i x assume "(\<lambda>i. u i x) ----> u' x" "\<forall>j. norm (u j x) \<le> w x" "\<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
 | 
1182  | 
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
 | 
1183  | 
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
 | 
1184  | 
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
 | 
1185  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1186  | 
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
 | 
1187  | 
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
 | 
1188  | 
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
 | 
1189  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1190  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1191  | 
have "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) ----> (\<integral>\<^sup>+x. 0 \<partial>M)"  | 
| 56996 | 1192  | 
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
 | 
1193  | 
show "(\<integral>\<^sup>+x. 2 * w x \<partial>M) < \<infinity>"  | 
| 56996 | 1194  | 
by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) auto  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1195  | 
show "AE x in M. (\<lambda>i. ereal (norm (u' x - u i x))) ----> 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1196  | 
using u'  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1197  | 
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
 | 
1198  | 
fix x assume "(\<lambda>i. u i x) ----> u' x"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1199  | 
from tendsto_diff[OF tendsto_const[of "u' x"] this]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1200  | 
show "(\<lambda>i. ereal (norm (u' x - u i x))) ----> 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1201  | 
by (simp add: zero_ereal_def tendsto_norm_zero_iff)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1202  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1203  | 
qed (insert bnd, auto)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1204  | 
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
 | 
1205  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1206  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1207  | 
lemma integrableI_bounded:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1208  | 
  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
 | 
1209  | 
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
 | 
1210  | 
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
 | 
1211  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1212  | 
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
 | 
1213  | 
s: "\<And>i. simple_function M (s i)" and  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1214  | 
pointwise: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) ----> f x" and  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1215  | 
bound: "\<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
 | 
1216  | 
by (simp add: norm_conv_dist) metis  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1217  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1218  | 
show ?thesis  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1219  | 
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
 | 
1220  | 
    { fix i
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1221  | 
have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. 2 * ereal (norm (f x)) \<partial>M)"  | 
| 56996 | 1222  | 
by (intro nn_integral_mono) (simp add: bound)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1223  | 
also have "\<dots> = 2 * (\<integral>\<^sup>+x. ereal (norm (f x)) \<partial>M)"  | 
| 56996 | 1224  | 
by (rule 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
 | 
1225  | 
finally have "(\<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
 | 
1226  | 
using fin by auto }  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1227  | 
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
 | 
1228  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1229  | 
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
 | 
1230  | 
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
 | 
1231  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1232  | 
show "(\<lambda>i. \<integral>\<^sup>+ x. ereal (norm (f x - s i x)) \<partial>M) ----> 0"  | 
| 56996 | 1233  | 
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
 | 
1234  | 
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
 | 
1235  | 
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
 | 
1236  | 
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
 | 
1237  | 
using s by (auto 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
 | 
1238  | 
show "(\<integral>\<^sup>+ x. ereal (2 * norm (f x)) \<partial>M) < \<infinity>"  | 
| 56996 | 1239  | 
using fin unfolding times_ereal.simps(1)[symmetric] by (subst 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
 | 
1240  | 
show "AE x in M. (\<lambda>i. s i 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
 | 
1241  | 
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
 | 
1242  | 
qed fact  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1243  | 
qed fact  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1244  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1245  | 
|
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1246  | 
lemma integrableI_bounded_set:  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1247  | 
  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
 | 
1248  | 
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
 | 
1249  | 
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
 | 
1250  | 
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
 | 
1251  | 
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
 | 
1252  | 
shows "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
 | 
1253  | 
proof (rule integrableI_bounded)  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1254  | 
  { 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
 | 
1255  | 
using norm_ge_zero[of x] by arith }  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1256  | 
with bnd null have "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (max 0 B) * indicator A 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
 | 
1257  | 
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
 | 
1258  | 
also have "\<dots> < \<infinity>"  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1259  | 
using finite by (subst nn_integral_cmult_indicator) (auto simp: max_def)  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1260  | 
finally show "(\<integral>\<^sup>+ x. ereal (norm (f 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
 | 
1261  | 
qed simp  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1262  | 
|
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1263  | 
lemma integrableI_bounded_set_indicator:  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1264  | 
  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
 | 
1265  | 
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
 | 
1266  | 
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
 | 
1267  | 
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
 | 
1268  | 
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
 | 
1269  | 
|
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1270  | 
lemma integrableI_nonneg:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1271  | 
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
 | 
1272  | 
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
 | 
1273  | 
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
 | 
1274  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1275  | 
have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"  | 
| 56996 | 1276  | 
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
 | 
1277  | 
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
 | 
1278  | 
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
 | 
1279  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1280  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1281  | 
lemma 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
 | 
1282  | 
  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
 | 
1283  | 
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
 | 
1284  | 
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
 | 
1285  | 
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
 | 
1286  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1287  | 
lemma integrable_bound:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1288  | 
  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
 | 
1289  | 
    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
 | 
1290  | 
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
 | 
1291  | 
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
 | 
1292  | 
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
 | 
1293  | 
proof safe  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1294  | 
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
 | 
1295  | 
assume "AE x in M. norm (g x) \<le> 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
 | 
1296  | 
then have "(\<integral>\<^sup>+ x. ereal (norm (g x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M)"  | 
| 56996 | 1297  | 
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
 | 
1298  | 
also assume "(\<integral>\<^sup>+ x. ereal (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
 | 
1299  | 
finally show "(\<integral>\<^sup>+ x. ereal (norm (g 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
 | 
1300  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1301  | 
|
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1302  | 
lemma integrable_mult_indicator:  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1303  | 
  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
 | 
1304  | 
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
 | 
1305  | 
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
 | 
1306  | 
|
| 59000 | 1307  | 
lemma integrable_real_mult_indicator:  | 
1308  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
1309  | 
shows "A \<in> sets M \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. f x * indicator A x)"  | 
|
1310  | 
using integrable_mult_indicator[of A M f] by (simp add: mult_ac)  | 
|
1311  | 
||
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1312  | 
lemma integrable_abs[simp, intro]:  | 
| 
 
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> real"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1314  | 
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
 | 
1315  | 
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
 | 
1316  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1317  | 
lemma integrable_norm[simp, intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1318  | 
  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
 | 
1319  | 
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
 | 
1320  | 
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
 | 
1321  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1322  | 
lemma integrable_norm_cancel:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1323  | 
  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
 | 
1324  | 
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
 | 
1325  | 
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
 | 
1326  | 
|
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1327  | 
lemma integrable_norm_iff:  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1328  | 
  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
 | 
1329  | 
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
 | 
1330  | 
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
 | 
1331  | 
|
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1332  | 
lemma integrable_abs_cancel:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1333  | 
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
 | 
1334  | 
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
 | 
1335  | 
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
 | 
1336  | 
|
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1337  | 
lemma integrable_abs_iff:  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1338  | 
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
 | 
1339  | 
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
 | 
1340  | 
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
 | 
1341  | 
|
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1342  | 
lemma integrable_max[simp, intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1343  | 
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
 | 
1344  | 
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
 | 
1345  | 
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
 | 
1346  | 
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
 | 
1347  | 
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
 | 
1348  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1349  | 
lemma integrable_min[simp, intro]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1350  | 
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
 | 
1351  | 
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
 | 
1352  | 
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
 | 
1353  | 
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
 | 
1354  | 
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
 | 
1355  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1356  | 
lemma integral_minus_iff[simp]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1357  | 
  "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
 | 
1358  | 
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
 | 
1359  | 
by (auto intro: borel_measurable_uminus[of "\<lambda>x. - f x" M, simplified])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1360  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1361  | 
lemma 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
 | 
1362  | 
"integrable M (indicator A::_ \<Rightarrow> real) \<longleftrightarrow> A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>"  | 
| 56996 | 1363  | 
by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ereal_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
 | 
1364  | 
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
 | 
1365  | 
|
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1366  | 
lemma integral_indicator[simp]: "integral\<^sup>L M (indicator A) = 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
 | 
1367  | 
proof cases  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1368  | 
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
 | 
1369  | 
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
 | 
1370  | 
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
 | 
1371  | 
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
 | 
1372  | 
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
 | 
1373  | 
finally show ?thesis .  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1374  | 
next  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1375  | 
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
 | 
1376  | 
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
 | 
1377  | 
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
 | 
1378  | 
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
 | 
1379  | 
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
 | 
1380  | 
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
 | 
1381  | 
using * by (auto simp: measure_def emeasure_notin_sets)  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1382  | 
finally show ?thesis .  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1383  | 
qed  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1384  | 
|
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1385  | 
lemma integrable_discrete_difference:  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1386  | 
  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
 | 
1387  | 
assumes X: "countable X"  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1388  | 
  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" 
 | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1389  | 
  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
 | 
1390  | 
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
 | 
1391  | 
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
 | 
1392  | 
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
 | 
1393  | 
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
 | 
1394  | 
  { 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
 | 
1395  | 
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
 | 
1396  | 
moreover  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1397  | 
  { 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
 | 
1398  | 
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
 | 
1399  | 
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
 | 
1400  | 
next  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1401  | 
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
 | 
1402  | 
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
 | 
1403  | 
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
 | 
1404  | 
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
 | 
1405  | 
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
 | 
1406  | 
by simp  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1407  | 
qed  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1408  | 
|
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1409  | 
lemma integral_discrete_difference:  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1410  | 
  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
 | 
1411  | 
assumes X: "countable X"  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1412  | 
  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" 
 | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1413  | 
  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
 | 
1414  | 
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
 | 
1415  | 
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
 | 
1416  | 
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
 | 
1417  | 
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
 | 
1418  | 
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
 | 
1419  | 
|
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1420  | 
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
 | 
1421  | 
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
 | 
1422  | 
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
 | 
1423  | 
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
 | 
1424  | 
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
 | 
1425  | 
|
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1426  | 
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
 | 
1427  | 
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
 | 
1428  | 
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
 | 
1429  | 
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
 | 
1430  | 
qed  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1431  | 
qed  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1432  | 
|
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1433  | 
lemma has_bochner_integral_discrete_difference:  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1434  | 
  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
 | 
1435  | 
assumes X: "countable X"  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1436  | 
  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" 
 | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1437  | 
  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
 | 
1438  | 
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
 | 
1439  | 
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
 | 
1440  | 
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
 | 
1441  | 
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
 | 
1442  | 
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
 | 
1443  | 
|
| 57137 | 1444  | 
lemma  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1445  | 
  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
 | 
1446  | 
assumes "f \<in> borel_measurable M" "\<And>i. s i \<in> borel_measurable M" "integrable M w"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1447  | 
assumes lim: "AE x in M. (\<lambda>i. s i 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
 | 
1448  | 
assumes bound: "\<And>i. AE x in M. norm (s i x) \<le> w x"  | 
| 57137 | 1449  | 
shows integrable_dominated_convergence: "integrable M f"  | 
1450  | 
and integrable_dominated_convergence2: "\<And>i. integrable M (s i)"  | 
|
1451  | 
and integral_dominated_convergence: "(\<lambda>i. integral\<^sup>L M (s i)) ----> 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
 | 
1452  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1453  | 
have "AE x in M. 0 \<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
 | 
1454  | 
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
 | 
1455  | 
then have "(\<integral>\<^sup>+x. w x \<partial>M) = (\<integral>\<^sup>+x. norm (w x) \<partial>M)"  | 
| 56996 | 1456  | 
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
 | 
1457  | 
with `integrable M w` have w: "w \<in> borel_measurable M" "(\<integral>\<^sup>+x. w 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
 | 
1458  | 
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
 | 
1459  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1460  | 
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
 | 
1461  | 
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
 | 
1462  | 
proof  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1463  | 
fix i  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1464  | 
have "(\<integral>\<^sup>+ x. ereal (norm (s i x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)"  | 
| 56996 | 1465  | 
using bound 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
 | 
1466  | 
with w show "(\<integral>\<^sup>+ x. ereal (norm (s i x)) \<partial>M) < \<infinity>" by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1467  | 
qed fact  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1468  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1469  | 
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
 | 
1470  | 
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
 | 
1471  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1472  | 
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
 | 
1473  | 
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
 | 
1474  | 
proof  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1475  | 
have "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w 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
 | 
1476  | 
using all_bound lim  | 
| 56996 | 1477  | 
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
 | 
1478  | 
fix x assume "\<forall>i. norm (s i x) \<le> w x" "(\<lambda>i. s i 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
 | 
1479  | 
then show "ereal (norm (f x)) \<le> ereal (w x)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1480  | 
by (intro LIMSEQ_le_const2[where X="\<lambda>i. ereal (norm (s i x))"] tendsto_intros lim_ereal[THEN iffD2]) auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1481  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1482  | 
with w show "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) < \<infinity>" by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1483  | 
qed fact  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1484  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1485  | 
have "(\<lambda>n. ereal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) ----> ereal 0" (is "?d ----> ereal 0")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1486  | 
proof (rule tendsto_sandwich)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1487  | 
show "eventually (\<lambda>n. ereal 0 \<le> ?d n) sequentially" "(\<lambda>_. ereal 0) ----> ereal 0" by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1488  | 
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
 | 
1489  | 
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
 | 
1490  | 
fix n  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1491  | 
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
 | 
1492  | 
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
 | 
1493  | 
also have "\<dots> \<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
 | 
1494  | 
by (intro int_f int_s integrable_diff integral_norm_bound_ereal)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1495  | 
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
 | 
1496  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1497  | 
show "(\<lambda>n. \<integral>\<^sup>+x. norm (s n x - f x) \<partial>M) ----> ereal 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1498  | 
unfolding zero_ereal_def[symmetric]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1499  | 
apply (subst norm_minus_commute)  | 
| 56996 | 1500  | 
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
 | 
1501  | 
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
 | 
1502  | 
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
 | 
1503  | 
qed fact+  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1504  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1505  | 
then have "(\<lambda>n. integral\<^sup>L M (s n) - integral\<^sup>L M f) ----> 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1506  | 
unfolding lim_ereal tendsto_norm_zero_iff .  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1507  | 
from tendsto_add[OF this tendsto_const[of "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
 | 
1508  | 
show "(\<lambda>i. integral\<^sup>L M (s i)) ----> integral\<^sup>L 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
 | 
1509  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1510  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1511  | 
lemma integrable_mult_left_iff:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1512  | 
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
 | 
1513  | 
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
 | 
1514  | 
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
 | 
1515  | 
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
 | 
1516  | 
|
| 56996 | 1517  | 
lemma 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
 | 
1518  | 
assumes 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
 | 
1519  | 
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
 | 
1520  | 
shows "(\<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
 | 
1521  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1522  | 
  { 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
 | 
1523  | 
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
 | 
1524  | 
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
 | 
1525  | 
case (set A) then show ?case  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1526  | 
by (simp add: integrable_indicator_iff ereal_indicator emeasure_eq_ereal_measure)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1527  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1528  | 
case (mult f c) then show ?case  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1529  | 
unfolding times_ereal.simps(1)[symmetric]  | 
| 56996 | 1530  | 
by (subst nn_integral_cmult)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1531  | 
(auto simp add: integrable_mult_left_iff zero_ereal_def[symmetric])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1532  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1533  | 
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
 | 
1534  | 
then have "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
 | 
1535  | 
by (auto intro!: integrable_bound[OF add(8)])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1536  | 
with add show ?case  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1537  | 
unfolding plus_ereal.simps(1)[symmetric]  | 
| 56996 | 1538  | 
by (subst 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
 | 
1539  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1540  | 
case (seq s)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1541  | 
      { fix i x assume "x \<in> space M" with seq(4) have "s i 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
 | 
1542  | 
by (intro LIMSEQ_le_const[OF seq(5)] exI[of _ i]) (auto simp: incseq_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
 | 
1543  | 
note s_le_f = this  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1544  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1545  | 
show ?case  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1546  | 
proof (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
 | 
1547  | 
show "(\<lambda>i. ereal (integral\<^sup>L M (s i))) ----> ereal (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
 | 
1548  | 
unfolding lim_ereal  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1549  | 
proof (rule integral_dominated_convergence[where w=f])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1550  | 
show "integrable M f" by fact  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1551  | 
from s_le_f seq show "\<And>i. AE x in M. norm (s i 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
 | 
1552  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1553  | 
qed (insert seq, auto)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1554  | 
have 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
 | 
1555  | 
using seq f s_le_f by (intro integrable_bound[OF f(3)]) auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1556  | 
have "(\<lambda>i. \<integral>\<^sup>+ x. s i x \<partial>M) ----> \<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
 | 
1557  | 
using seq s_le_f f  | 
| 56996 | 1558  | 
by (intro nn_integral_dominated_convergence[where w=f])  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1559  | 
(auto simp: 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
 | 
1560  | 
also have "(\<lambda>i. \<integral>\<^sup>+x. s i x \<partial>M) = (\<lambda>i. \<integral>x. s i 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
 | 
1561  | 
using seq 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
 | 
1562  | 
finally show "(\<lambda>i. \<integral>x. s i x \<partial>M) ----> \<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
 | 
1563  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1564  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1565  | 
qed }  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1566  | 
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
 | 
1567  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1568  | 
also have "\<dots> = 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
 | 
1569  | 
using assms by (auto intro!: 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
 | 
1570  | 
also have "(\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)"  | 
| 56996 | 1571  | 
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
 | 
1572  | 
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
 | 
1573  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1574  | 
|
| 57036 | 1575  | 
lemma  | 
1576  | 
  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> 'a :: {banach, second_countable_topology}"
 | 
|
1577  | 
assumes integrable[measurable]: "\<And>i. integrable M (f i)"  | 
|
1578  | 
and summable: "AE x in M. summable (\<lambda>i. norm (f i x))"  | 
|
1579  | 
and sums: "summable (\<lambda>i. (\<integral>x. norm (f i x) \<partial>M))"  | 
|
1580  | 
shows integrable_suminf: "integrable M (\<lambda>x. (\<Sum>i. f i x))" (is "integrable M ?S")  | 
|
1581  | 
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")  | 
|
1582  | 
and integral_suminf: "(\<integral>x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>L M (f i))"  | 
|
1583  | 
and summable_integral: "summable (\<lambda>i. integral\<^sup>L M (f i))"  | 
|
1584  | 
proof -  | 
|
1585  | 
have 1: "integrable M (\<lambda>x. \<Sum>i. norm (f i x))"  | 
|
1586  | 
proof (rule integrableI_bounded)  | 
|
1587  | 
have "(\<integral>\<^sup>+ x. ereal (norm (\<Sum>i. norm (f i x))) \<partial>M) = (\<integral>\<^sup>+ x. (\<Sum>i. ereal (norm (f i x))) \<partial>M)"  | 
|
1588  | 
apply (intro nn_integral_cong_AE)  | 
|
1589  | 
using summable  | 
|
1590  | 
apply eventually_elim  | 
|
1591  | 
apply (simp add: suminf_ereal' suminf_nonneg)  | 
|
1592  | 
done  | 
|
1593  | 
also have "\<dots> = (\<Sum>i. \<integral>\<^sup>+ x. norm (f i x) \<partial>M)"  | 
|
1594  | 
by (intro nn_integral_suminf) auto  | 
|
1595  | 
also have "\<dots> = (\<Sum>i. ereal (\<integral>x. norm (f i x) \<partial>M))"  | 
|
1596  | 
by (intro arg_cong[where f=suminf] ext nn_integral_eq_integral integrable_norm integrable) auto  | 
|
1597  | 
finally show "(\<integral>\<^sup>+ x. ereal (norm (\<Sum>i. norm (f i x))) \<partial>M) < \<infinity>"  | 
|
1598  | 
by (simp add: suminf_ereal' sums)  | 
|
1599  | 
qed simp  | 
|
1600  | 
||
1601  | 
have 2: "AE x in M. (\<lambda>n. \<Sum>i<n. f i x) ----> (\<Sum>i. f i x)"  | 
|
1602  | 
using summable by eventually_elim (auto intro: summable_LIMSEQ summable_norm_cancel)  | 
|
1603  | 
||
1604  | 
have 3: "\<And>j. AE x in M. norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))"  | 
|
1605  | 
using summable  | 
|
1606  | 
proof eventually_elim  | 
|
1607  | 
fix j x assume [simp]: "summable (\<lambda>i. norm (f i x))"  | 
|
1608  | 
have "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i<j. norm (f i x))" by (rule norm_setsum)  | 
|
1609  | 
also have "\<dots> \<le> (\<Sum>i. norm (f i x))"  | 
|
1610  | 
using setsum_le_suminf[of "\<lambda>i. norm (f i x)"] unfolding sums_iff by auto  | 
|
1611  | 
finally show "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))" by simp  | 
|
1612  | 
qed  | 
|
1613  | 
||
| 57137 | 1614  | 
note ibl = integrable_dominated_convergence[OF _ _ 1 2 3]  | 
1615  | 
note int = integral_dominated_convergence[OF _ _ 1 2 3]  | 
|
| 57036 | 1616  | 
|
1617  | 
show "integrable M ?S"  | 
|
| 57137 | 1618  | 
by (rule ibl) measurable  | 
| 57036 | 1619  | 
|
1620  | 
show "?f sums ?x" unfolding sums_def  | 
|
| 57137 | 1621  | 
using int by (simp add: integrable)  | 
| 57036 | 1622  | 
then show "?x = suminf ?f" "summable ?f"  | 
1623  | 
unfolding sums_iff by auto  | 
|
1624  | 
qed  | 
|
1625  | 
||
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1626  | 
lemma 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
 | 
1627  | 
  fixes f :: "_ \<Rightarrow> 'a :: {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
 | 
1628  | 
shows "integrable M f \<Longrightarrow> norm (integral\<^sup>L M f) \<le> (\<integral>x. norm (f x) \<partial>M)"  | 
| 56996 | 1629  | 
using nn_integral_eq_integral[of M "\<lambda>x. 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
 | 
1630  | 
using integral_norm_bound_ereal[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
 | 
1631  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1632  | 
lemma integrableI_nn_integral_finite:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1633  | 
assumes [measurable]: "f \<in> borel_measurable M"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1634  | 
and nonneg: "AE x in M. 0 \<le> f x"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1635  | 
and finite: "(\<integral>\<^sup>+x. f x \<partial>M) = ereal x"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1636  | 
shows "integrable M f"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1637  | 
proof (rule integrableI_bounded)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1638  | 
have "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) = (\<integral>\<^sup>+ x. ereal (f x) \<partial>M)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1639  | 
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
 | 
1640  | 
with finite show "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) < \<infinity>"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1641  | 
by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1642  | 
qed simp  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1643  | 
|
| 56996 | 1644  | 
lemma integral_eq_nn_integral:  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1645  | 
assumes [measurable]: "f \<in> borel_measurable M"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1646  | 
assumes nonneg: "AE x in M. 0 \<le> f x"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1647  | 
shows "integral\<^sup>L M f = real (\<integral>\<^sup>+ x. ereal (f x) \<partial>M)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1648  | 
proof cases  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1649  | 
assume *: "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) = \<infinity>"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1650  | 
also have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) = (\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1651  | 
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
 | 
1652  | 
finally have "\<not> integrable M f"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1653  | 
by (auto simp: integrable_iff_bounded)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1654  | 
then show ?thesis  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1655  | 
by (simp add: * not_integrable_integral_eq)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1656  | 
next  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1657  | 
assume "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1658  | 
then have "integrable M f"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1659  | 
by (cases "\<integral>\<^sup>+ x. ereal (f x) \<partial>M") (auto intro!: integrableI_nn_integral_finite assms)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1660  | 
from nn_integral_eq_integral[OF this nonneg] show ?thesis  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1661  | 
by simp  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
1662  | 
qed  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1663  | 
|
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1664  | 
lemma has_bochner_integral_nn_integral:  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1665  | 
assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> 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
 | 
1666  | 
assumes "(\<integral>\<^sup>+x. f x \<partial>M) = ereal x"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1667  | 
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
 | 
1668  | 
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
 | 
1669  | 
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
 | 
1670  | 
|
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1671  | 
lemma 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
 | 
1672  | 
  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
 | 
1673  | 
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
 | 
1674  | 
by (intro integrableI_sequence[where s="\<lambda>_. f"] 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
 | 
1675  | 
(auto simp: zero_ereal_def[symmetric] 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
 | 
1676  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1677  | 
lemma integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1678  | 
  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
 | 
1679  | 
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
 | 
1680  | 
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
 | 
1681  | 
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
 | 
1682  | 
assumes lim: "\<And>f s. (\<And>i. integrable M (s i)) \<Longrightarrow> (\<And>i. P (s i)) \<Longrightarrow>  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1683  | 
(\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) ----> 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
 | 
1684  | 
(\<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
 | 
1685  | 
shows "P f"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1686  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1687  | 
from `integrable M f` have f: "f \<in> borel_measurable M" "(\<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
 | 
1688  | 
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
 | 
1689  | 
from borel_measurable_implies_sequence_metric[OF f(1)]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1690  | 
obtain s where s: "\<And>i. simple_function M (s i)" "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i 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
 | 
1691  | 
"\<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
 | 
1692  | 
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
 | 
1693  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1694  | 
  { fix f A 
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1695  | 
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
 | 
1696  | 
      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
 | 
1697  | 
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
 | 
1698  | 
(\<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
 | 
1699  | 
by (induct A rule: infinite_finite_induct) (auto intro!: add) }  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1700  | 
note setsum = this  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1701  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1702  | 
def s' \<equiv> "\<lambda>i z. indicator (space M) z *\<^sub>R s i z"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1703  | 
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
 | 
1704  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1705  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1706  | 
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
 | 
1707  | 
unfolding s'_def using s(1)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1708  | 
by (intro simple_function_compose2[where h="op *\<^sub>R"] simple_function_indicator) auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1709  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1710  | 
  { fix i 
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1711  | 
    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
 | 
1712  | 
        (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
 | 
1713  | 
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
 | 
1714  | 
    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
 | 
1715  | 
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
 | 
1716  | 
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
 | 
1717  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1718  | 
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
 | 
1719  | 
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
 | 
1720  | 
fix i  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1721  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1722  | 
have "(\<integral>\<^sup>+x. norm (s' i x) \<partial>M) \<le> (\<integral>\<^sup>+x. 2 * ereal (norm (f x)) \<partial>M)"  | 
| 56996 | 1723  | 
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
 | 
1724  | 
also have "\<dots> < \<infinity>"  | 
| 56996 | 1725  | 
using f by (subst 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
 | 
1726  | 
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
 | 
1727  | 
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
 | 
1728  | 
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
 | 
1729  | 
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
 | 
1730  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1731  | 
    { 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
 | 
1732  | 
      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
 | 
1733  | 
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
 | 
1734  | 
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
 | 
1735  | 
using sbi by (auto 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
 | 
1736  | 
      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
 | 
1737  | 
then show "P (s' i)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1738  | 
by (subst s'_eq) (auto intro!: setsum base)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1739  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1740  | 
fix x assume "x \<in> space M" with s show "(\<lambda>i. s' i 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
 | 
1741  | 
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
 | 
1742  | 
show "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
 | 
1743  | 
using `x \<in> space M` s 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
 | 
1744  | 
qed fact  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1745  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1746  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1747  | 
lemma integral_nonneg_AE:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1748  | 
fixes f :: "'a \<Rightarrow> real"  | 
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1749  | 
assumes [measurable]: "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
 | 
1750  | 
shows "0 \<le> 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
 | 
1751  | 
proof cases  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1752  | 
assume [measurable]: "integrable M f"  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1753  | 
then have "0 \<le> ereal (integral\<^sup>L M (\<lambda>x. max 0 (f x)))"  | 
| 57036 | 1754  | 
by (subst integral_eq_nn_integral) (auto intro: real_of_ereal_pos nn_integral_nonneg assms)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1755  | 
also have "integral\<^sup>L M (\<lambda>x. max 0 (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
 | 
1756  | 
using assms by (intro integral_cong_AE assms integrable_max) auto  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1757  | 
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
 | 
1758  | 
by simp  | 
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1759  | 
qed (simp 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
 | 
1760  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1761  | 
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
 | 
1762  | 
"(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
 | 
1763  | 
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
 | 
1764  | 
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
 | 
1765  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1766  | 
lemma integral_nonneg_eq_0_iff_AE:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1767  | 
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
 | 
1768  | 
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
 | 
1769  | 
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
 | 
1770  | 
proof  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1771  | 
assume "integral\<^sup>L M f = 0"  | 
| 56996 | 1772  | 
then have "integral\<^sup>N M f = 0"  | 
1773  | 
using nn_integral_eq_integral[OF f nonneg] 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
 | 
1774  | 
then have "AE x in M. ereal (f x) \<le> 0"  | 
| 56996 | 1775  | 
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
 | 
1776  | 
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
 | 
1777  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1778  | 
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
 | 
1779  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1780  | 
lemma integral_mono_AE:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1781  | 
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
 | 
1782  | 
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
 | 
1783  | 
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
 | 
1784  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1785  | 
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
 | 
1786  | 
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
 | 
1787  | 
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
 | 
1788  | 
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
 | 
1789  | 
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
 | 
1790  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1791  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1792  | 
lemma integral_mono:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1793  | 
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
 | 
1794  | 
shows "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x \<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
 | 
1795  | 
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
 | 
1796  | 
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
 | 
1797  | 
|
| 57025 | 1798  | 
lemma (in finite_measure) integrable_measure:  | 
1799  | 
assumes I: "disjoint_family_on X I" "countable I"  | 
|
1800  | 
shows "integrable (count_space I) (\<lambda>i. measure M (X i))"  | 
|
1801  | 
proof -  | 
|
1802  | 
  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)"
 | 
|
1803  | 
by (auto intro!: nn_integral_cong measure_notin_sets)  | 
|
1804  | 
  also have "\<dots> = measure M (\<Union>i\<in>I. if X i \<in> sets M then X i else {})"
 | 
|
1805  | 
using I unfolding emeasure_eq_measure[symmetric]  | 
|
1806  | 
by (subst emeasure_UN_countable) (auto simp: disjoint_family_on_def)  | 
|
1807  | 
finally show ?thesis  | 
|
1808  | 
by (auto intro!: integrableI_bounded simp: measure_nonneg)  | 
|
1809  | 
qed  | 
|
1810  | 
||
1811  | 
lemma integrableI_real_bounded:  | 
|
1812  | 
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>"  | 
|
1813  | 
shows "integrable M f"  | 
|
1814  | 
proof (rule integrableI_bounded)  | 
|
1815  | 
have "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) = \<integral>\<^sup>+ x. ereal (f x) \<partial>M"  | 
|
1816  | 
using ae by (auto intro: nn_integral_cong_AE)  | 
|
1817  | 
also note fin  | 
|
1818  | 
finally show "(\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M) < \<infinity>" .  | 
|
1819  | 
qed fact  | 
|
1820  | 
||
1821  | 
lemma integral_real_bounded:  | 
|
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1822  | 
assumes "AE x in M. 0 \<le> f x" "integral\<^sup>N M f \<le> ereal r"  | 
| 57025 | 1823  | 
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
 | 
1824  | 
proof cases  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1825  | 
assume "integrable M f" from nn_integral_eq_integral[OF this] assms show ?thesis  | 
| 57025 | 1826  | 
by simp  | 
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1827  | 
next  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1828  | 
assume "\<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
 | 
1829  | 
moreover have "0 \<le> ereal r"  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1830  | 
using nn_integral_nonneg assms(2) by (rule order_trans)  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1831  | 
ultimately show ?thesis  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
1832  | 
by (simp add: not_integrable_integral_eq)  | 
| 57025 | 1833  | 
qed  | 
1834  | 
||
| 57137 | 1835  | 
subsection {* Restricted measure spaces *}
 | 
1836  | 
||
1837  | 
lemma integrable_restrict_space:  | 
|
1838  | 
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
|
1839  | 
assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"  | 
|
1840  | 
shows "integrable (restrict_space M \<Omega>) f \<longleftrightarrow> integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"  | 
|
1841  | 
unfolding integrable_iff_bounded  | 
|
1842  | 
borel_measurable_restrict_space_iff[OF \<Omega>]  | 
|
1843  | 
nn_integral_restrict_space[OF \<Omega>]  | 
|
1844  | 
by (simp add: ac_simps ereal_indicator times_ereal.simps(1)[symmetric] del: times_ereal.simps)  | 
|
1845  | 
||
1846  | 
lemma integral_restrict_space:  | 
|
1847  | 
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
|
1848  | 
assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"  | 
|
1849  | 
shows "integral\<^sup>L (restrict_space M \<Omega>) f = integral\<^sup>L M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"  | 
|
1850  | 
proof (rule integral_eq_cases)  | 
|
1851  | 
assume "integrable (restrict_space M \<Omega>) f"  | 
|
1852  | 
then show ?thesis  | 
|
1853  | 
proof induct  | 
|
1854  | 
case (base A c) then show ?case  | 
|
1855  | 
by (simp add: indicator_inter_arith[symmetric] sets_restrict_space_iff  | 
|
1856  | 
emeasure_restrict_space Int_absorb1 measure_restrict_space)  | 
|
1857  | 
next  | 
|
1858  | 
case (add g f) then show ?case  | 
|
1859  | 
by (simp add: scaleR_add_right integrable_restrict_space)  | 
|
1860  | 
next  | 
|
1861  | 
case (lim f s)  | 
|
1862  | 
show ?case  | 
|
1863  | 
proof (rule LIMSEQ_unique)  | 
|
1864  | 
show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) ----> integral\<^sup>L (restrict_space M \<Omega>) f"  | 
|
1865  | 
using lim by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) simp_all  | 
|
1866  | 
||
1867  | 
show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) ----> (\<integral> x. indicator \<Omega> x *\<^sub>R f x \<partial>M)"  | 
|
1868  | 
unfolding lim  | 
|
1869  | 
using lim  | 
|
1870  | 
by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (indicator \<Omega> x *\<^sub>R f x)"])  | 
|
1871  | 
(auto simp add: space_restrict_space integrable_restrict_space  | 
|
1872  | 
simp del: norm_scaleR  | 
|
1873  | 
split: split_indicator)  | 
|
1874  | 
qed  | 
|
1875  | 
qed  | 
|
1876  | 
qed (simp add: integrable_restrict_space)  | 
|
1877  | 
||
| 60066 | 1878  | 
lemma integral_empty:  | 
1879  | 
  assumes "space M = {}"
 | 
|
1880  | 
shows "integral\<^sup>L M f = 0"  | 
|
1881  | 
proof -  | 
|
1882  | 
have "(\<integral> x. f x \<partial>M) = (\<integral> x. 0 \<partial>M)"  | 
|
1883  | 
by(rule integral_cong)(simp_all add: assms)  | 
|
1884  | 
thus ?thesis by simp  | 
|
1885  | 
qed  | 
|
1886  | 
||
| 56994 | 1887  | 
subsection {* Measure spaces with an associated density *}
 | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1888  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1889  | 
lemma integrable_density:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1890  | 
  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
 | 
1891  | 
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
 | 
1892  | 
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
 | 
1893  | 
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
 | 
1894  | 
unfolding integrable_iff_bounded using nn  | 
| 56996 | 1895  | 
apply (simp add: nn_integral_density )  | 
1896  | 
apply (intro arg_cong2[where f="op ="] refl 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
 | 
1897  | 
apply auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1898  | 
done  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1899  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1900  | 
lemma integral_density:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1901  | 
  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
 | 
1902  | 
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
 | 
1903  | 
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
 | 
1904  | 
shows "integral\<^sup>L (density M g) f = integral\<^sup>L 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
 | 
1905  | 
proof (rule integral_eq_cases)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1906  | 
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
 | 
1907  | 
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
 | 
1908  | 
proof induct  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1909  | 
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
 | 
1910  | 
then have [measurable]: "A \<in> sets 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
 | 
1911  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1912  | 
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
 | 
1913  | 
using g base integrable_density[of "indicator A :: 'a \<Rightarrow> real" M 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
 | 
1914  | 
then have "integral\<^sup>L M (\<lambda>x. g x * indicator A x) = (\<integral>\<^sup>+ x. ereal (g x * indicator A x) \<partial>M)"  | 
| 56996 | 1915  | 
using g by (subst 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
 | 
1916  | 
also have "\<dots> = (\<integral>\<^sup>+ x. ereal (g x) * indicator A x \<partial>M)"  | 
| 56996 | 1917  | 
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
 | 
1918  | 
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
 | 
1919  | 
by (rule emeasure_density[symmetric]) 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> = ereal (measure (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
 | 
1921  | 
using base by (auto intro: emeasure_eq_ereal_measure)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1922  | 
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
 | 
1923  | 
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
 | 
1924  | 
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
 | 
1925  | 
using base by (simp add: int)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1926  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1927  | 
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
 | 
1928  | 
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
 | 
1929  | 
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
 | 
1930  | 
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
 | 
1931  | 
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
 | 
1932  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1933  | 
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
 | 
1934  | 
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
 | 
1935  | 
using lim(1,5)[THEN borel_measurable_integrable] by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1936  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1937  | 
show ?case  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1938  | 
proof (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
 | 
1939  | 
show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) ----> integral\<^sup>L M (\<lambda>x. g x *\<^sub>R f x)"  | 
| 57137 | 1940  | 
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
 | 
1941  | 
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
 | 
1942  | 
by (intro integrable_mult_right integrable_norm integrable_density[THEN iffD1] lim g) auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1943  | 
show "AE x in M. (\<lambda>i. g x *\<^sub>R s i 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
 | 
1944  | 
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
 | 
1945  | 
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
 | 
1946  | 
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
 | 
1947  | 
qed auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1948  | 
show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) ----> integral\<^sup>L (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
 | 
1949  | 
unfolding lim(2)[symmetric]  | 
| 57137 | 1950  | 
by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])  | 
| 57036 | 1951  | 
(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
 | 
1952  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1953  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1954  | 
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
 | 
1955  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1956  | 
lemma  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1957  | 
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
 | 
1958  | 
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
 | 
1959  | 
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
 | 
1960  | 
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
 | 
1961  | 
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
 | 
1962  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1963  | 
lemma has_bochner_integral_density:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1964  | 
  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
 | 
1965  | 
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
 | 
1966  | 
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
 | 
1967  | 
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
 | 
1968  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1969  | 
subsection {* Distributions *}
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1970  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1971  | 
lemma 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
 | 
1972  | 
  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
 | 
1973  | 
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
 | 
1974  | 
shows "integrable (distr M N g) f \<longleftrightarrow> integrable M (\<lambda>x. f (g x))"  | 
| 56996 | 1975  | 
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
 | 
1976  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1977  | 
lemma integrable_distr:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1978  | 
  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
 | 
1979  | 
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
 | 
1980  | 
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
 | 
1981  | 
(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
 | 
1982  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1983  | 
lemma integral_distr:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1984  | 
  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
 | 
1985  | 
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
 | 
1986  | 
shows "integral\<^sup>L (distr M N g) f = integral\<^sup>L 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
 | 
1987  | 
proof (rule integral_eq_cases)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1988  | 
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
 | 
1989  | 
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
 | 
1990  | 
proof induct  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1991  | 
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
 | 
1992  | 
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
 | 
1993  | 
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
 | 
1994  | 
by (intro integrable_indicator)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1995  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
1996  | 
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
 | 
1997  | 
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
 | 
1998  | 
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
 | 
1999  | 
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
 | 
2000  | 
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
 | 
2001  | 
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
 | 
2002  | 
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
 | 
2003  | 
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
 | 
2004  | 
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
 | 
2005  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2006  | 
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
 | 
2007  | 
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
 | 
2008  | 
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
 | 
2009  | 
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
 | 
2010  | 
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
 | 
2011  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2012  | 
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
 | 
2013  | 
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
 | 
2014  | 
using lim(1,5)[THEN borel_measurable_integrable] by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2015  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2016  | 
show ?case  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2017  | 
proof (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
 | 
2018  | 
show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) ----> integral\<^sup>L M (\<lambda>x. f (g x))"  | 
| 57137 | 2019  | 
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
 | 
2020  | 
show "integrable M (\<lambda>x. 2 * norm (f (g x)))"  | 
| 57036 | 2021  | 
using lim by (auto simp: 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
 | 
2022  | 
show "AE x in M. (\<lambda>i. s i (g 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
 | 
2023  | 
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
 | 
2024  | 
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
 | 
2025  | 
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
 | 
2026  | 
qed auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2027  | 
show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) ----> integral\<^sup>L (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
 | 
2028  | 
unfolding lim(2)[symmetric]  | 
| 57137 | 2029  | 
by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])  | 
| 57036 | 2030  | 
(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
 | 
2031  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2032  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2033  | 
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
 | 
2034  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2035  | 
lemma has_bochner_integral_distr:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2036  | 
  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
 | 
2037  | 
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
 | 
2038  | 
has_bochner_integral M (\<lambda>x. f (g x)) x \<Longrightarrow> has_bochner_integral (distr M N 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
 | 
2039  | 
by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2040  | 
|
| 56994 | 2041  | 
subsection {* Lebesgue integration on @{const 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
 | 
2042  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2043  | 
lemma 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
 | 
2044  | 
  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
 | 
2045  | 
shows "finite X \<Longrightarrow> integrable (count_space X) f"  | 
| 56996 | 2046  | 
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
 | 
2047  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2048  | 
lemma measure_count_space[simp]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2049  | 
"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
 | 
2050  | 
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
 | 
2051  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2052  | 
lemma lebesgue_integral_count_space_finite_support:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2053  | 
  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
 | 
2054  | 
shows "(\<integral>x. f x \<partial>count_space A) = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. f a)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2055  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2056  | 
  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)"
 | 
| 57418 | 2057  | 
by (intro setsum.mono_neutral_cong_left) auto  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2058  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2059  | 
  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
 | 
2060  | 
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
 | 
2061  | 
  also have "\<dots> = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. measure (count_space A) {a} *\<^sub>R f a)"
 | 
| 57418 | 2062  | 
by (subst integral_setsum) (auto intro!: setsum.cong)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2063  | 
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
 | 
2064  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2065  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2066  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2067  | 
lemma lebesgue_integral_count_space_finite: "finite A \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2068  | 
by (subst lebesgue_integral_count_space_finite_support)  | 
| 57418 | 2069  | 
(auto intro!: setsum.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
 | 
2070  | 
|
| 59425 | 2071  | 
lemma integrable_count_space_nat_iff:  | 
2072  | 
  fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
 | 
|
2073  | 
shows "integrable (count_space UNIV) f \<longleftrightarrow> summable (\<lambda>x. norm (f x))"  | 
|
2074  | 
by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat summable_ereal suminf_ereal_finite)  | 
|
2075  | 
||
2076  | 
lemma sums_integral_count_space_nat:  | 
|
2077  | 
  fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
 | 
|
2078  | 
assumes *: "integrable (count_space UNIV) f"  | 
|
2079  | 
shows "f sums (integral\<^sup>L (count_space UNIV) f)"  | 
|
2080  | 
proof -  | 
|
2081  | 
  let ?f = "\<lambda>n i. indicator {n} i *\<^sub>R f i"
 | 
|
2082  | 
  have f': "\<And>n i. ?f n i = indicator {n} i *\<^sub>R f n"
 | 
|
2083  | 
by (auto simp: fun_eq_iff split: split_indicator)  | 
|
2084  | 
||
2085  | 
have "(\<lambda>i. \<integral>n. ?f i n \<partial>count_space UNIV) sums \<integral> n. (\<Sum>i. ?f i n) \<partial>count_space UNIV"  | 
|
2086  | 
proof (rule sums_integral)  | 
|
2087  | 
show "\<And>i. integrable (count_space UNIV) (?f i)"  | 
|
2088  | 
using * by (intro integrable_mult_indicator) auto  | 
|
2089  | 
show "AE n in count_space UNIV. summable (\<lambda>i. norm (?f i n))"  | 
|
2090  | 
      using summable_finite[of "{n}" "\<lambda>i. norm (?f i n)" for n] by simp
 | 
|
2091  | 
show "summable (\<lambda>i. \<integral> n. norm (?f i n) \<partial>count_space UNIV)"  | 
|
2092  | 
using * by (subst f') (simp add: integrable_count_space_nat_iff)  | 
|
2093  | 
qed  | 
|
2094  | 
also have "(\<integral> n. (\<Sum>i. ?f i n) \<partial>count_space UNIV) = (\<integral>n. f n \<partial>count_space UNIV)"  | 
|
2095  | 
    using suminf_finite[of "{n}" "\<lambda>i. ?f i n" for n] by (auto intro!: integral_cong)
 | 
|
2096  | 
also have "(\<lambda>i. \<integral>n. ?f i n \<partial>count_space UNIV) = f"  | 
|
2097  | 
by (subst f') simp  | 
|
2098  | 
finally show ?thesis .  | 
|
2099  | 
qed  | 
|
2100  | 
||
2101  | 
lemma integral_count_space_nat:  | 
|
2102  | 
  fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
 | 
|
2103  | 
shows "integrable (count_space UNIV) f \<Longrightarrow> integral\<^sup>L (count_space UNIV) f = (\<Sum>x. f x)"  | 
|
2104  | 
using sums_integral_count_space_nat by (rule sums_unique)  | 
|
2105  | 
||
| 56994 | 2106  | 
subsection {* Point measure *}
 | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2107  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2108  | 
lemma lebesgue_integral_point_measure_finite:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2109  | 
  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
 | 
2110  | 
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
 | 
2111  | 
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
 | 
2112  | 
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
 | 
2113  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2114  | 
lemma integrable_point_measure_finite:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2115  | 
  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
 | 
2116  | 
shows "finite A \<Longrightarrow> integrable (point_measure A f) g"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2117  | 
unfolding 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
 | 
2118  | 
apply (subst density_ereal_max_0)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2119  | 
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
 | 
2120  | 
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
 | 
2121  | 
done  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2122  | 
|
| 59425 | 2123  | 
subsection {* Lebesgue integration on @{const null_measure} *}
 | 
2124  | 
||
2125  | 
lemma has_bochner_integral_null_measure_iff[iff]:  | 
|
2126  | 
"has_bochner_integral (null_measure M) f 0 \<longleftrightarrow> f \<in> borel_measurable M"  | 
|
2127  | 
by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def]  | 
|
2128  | 
intro!: exI[of _ "\<lambda>n x. 0"] simple_bochner_integrable.intros)  | 
|
2129  | 
||
2130  | 
lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \<longleftrightarrow> f \<in> borel_measurable M"  | 
|
2131  | 
by (auto simp add: integrable.simps)  | 
|
2132  | 
||
2133  | 
lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0"  | 
|
2134  | 
by (cases "integrable (null_measure M) f")  | 
|
2135  | 
(auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq)  | 
|
2136  | 
||
| 56994 | 2137  | 
subsection {* Legacy lemmas for the real-valued Lebesgue integral *}
 | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2138  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2139  | 
lemma real_lebesgue_integral_def:  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
2140  | 
assumes f[measurable]: "integrable 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
 | 
2141  | 
shows "integral\<^sup>L M f = real (\<integral>\<^sup>+x. f x \<partial>M) - real (\<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
 | 
2142  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2143  | 
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
 | 
2144  | 
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
 | 
2145  | 
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
 | 
2146  | 
by (intro integral_diff integrable_max integrable_minus integrable_zero f)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2147  | 
also have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = real (\<integral>\<^sup>+x. max 0 (f x) \<partial>M)"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
2148  | 
by (subst integral_eq_nn_integral[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
 | 
2149  | 
also have "integral\<^sup>L M (\<lambda>x. max 0 (- f x)) = real (\<integral>\<^sup>+x. max 0 (- f x) \<partial>M)"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57166 
diff
changeset
 | 
2150  | 
by (subst integral_eq_nn_integral[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
 | 
2151  | 
also have "(\<lambda>x. ereal (max 0 (f x))) = (\<lambda>x. max 0 (ereal (f x)))"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2152  | 
by (auto simp: max_def)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2153  | 
also have "(\<lambda>x. ereal (max 0 (- f x))) = (\<lambda>x. max 0 (- ereal (f x)))"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2154  | 
by (auto simp: max_def)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2155  | 
finally show ?thesis  | 
| 56996 | 2156  | 
unfolding nn_integral_max_0 .  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2157  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2158  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2159  | 
lemma real_integrable_def:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2160  | 
"integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2161  | 
(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2162  | 
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
 | 
2163  | 
proof (safe del: notI)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2164  | 
assume *: "(\<integral>\<^sup>+ x. ereal (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
 | 
2165  | 
have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M)"  | 
| 56996 | 2166  | 
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
 | 
2167  | 
also note *  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2168  | 
finally show "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2169  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2170  | 
have "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (norm (f x)) \<partial>M)"  | 
| 56996 | 2171  | 
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
 | 
2172  | 
also note *  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2173  | 
finally show "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2174  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2175  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2176  | 
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
 | 
2177  | 
assume fin: "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2178  | 
have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. max 0 (ereal (f x)) + max 0 (ereal (- f x)) \<partial>M)"  | 
| 56996 | 2179  | 
by (intro nn_integral_cong) (auto 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
 | 
2180  | 
also have"\<dots> = (\<integral>\<^sup>+ x. max 0 (ereal (f x)) \<partial>M) + (\<integral>\<^sup>+ x. max 0 (ereal (- f x)) \<partial>M)"  | 
| 56996 | 2181  | 
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
 | 
2182  | 
also have "\<dots> < \<infinity>"  | 
| 56996 | 2183  | 
using fin by (auto simp: nn_integral_max_0)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2184  | 
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
 | 
2185  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2186  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2187  | 
lemma integrableD[dest]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2188  | 
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
 | 
2189  | 
shows "f \<in> borel_measurable M" "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2190  | 
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
 | 
2191  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2192  | 
lemma integrableE:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2193  | 
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
 | 
2194  | 
obtains r q where  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2195  | 
"(\<integral>\<^sup>+x. ereal (f x)\<partial>M) = ereal r"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2196  | 
"(\<integral>\<^sup>+x. ereal (-f x)\<partial>M) = ereal q"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2197  | 
"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
 | 
2198  | 
using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms]  | 
| 56996 | 2199  | 
using nn_integral_nonneg[of M "\<lambda>x. ereal (f x)"]  | 
2200  | 
using nn_integral_nonneg[of M "\<lambda>x. ereal (-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
 | 
2201  | 
by (cases rule: ereal2_cases[of "(\<integral>\<^sup>+x. ereal (-f x)\<partial>M)" "(\<integral>\<^sup>+x. ereal (f x)\<partial>M)"]) auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2202  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2203  | 
lemma integral_monotone_convergence_nonneg:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2204  | 
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
 | 
2205  | 
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
 | 
2206  | 
and pos: "\<And>i. AE x in M. 0 \<le> 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
 | 
2207  | 
and lim: "AE x in M. (\<lambda>i. f i x) ----> u x"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2208  | 
and ilim: "(\<lambda>i. integral\<^sup>L M (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
 | 
2209  | 
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
 | 
2210  | 
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
 | 
2211  | 
and "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
 | 
2212  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2213  | 
have "(\<integral>\<^sup>+ x. ereal (u x) \<partial>M) = (SUP n. (\<integral>\<^sup>+ x. ereal (f n x) \<partial>M))"  | 
| 56996 | 2214  | 
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
 | 
2215  | 
fix i  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2216  | 
from mono pos show "AE x in M. ereal (f i x) \<le> ereal (f (Suc i) x) \<and> 0 \<le> ereal (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
 | 
2217  | 
by eventually_elim (auto simp: mono_def)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2218  | 
show "(\<lambda>x. ereal (f i 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
 | 
2219  | 
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
 | 
2220  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2221  | 
show "(\<integral>\<^sup>+ x. ereal (u x) \<partial>M) = \<integral>\<^sup>+ x. (SUP i. ereal (f i x)) \<partial>M"  | 
| 56996 | 2222  | 
apply (rule 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
 | 
2223  | 
using lim mono  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2224  | 
by eventually_elim (simp add: SUP_eq_LIMSEQ[THEN iffD2])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2225  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2226  | 
also have "\<dots> = ereal x"  | 
| 56996 | 2227  | 
using mono i unfolding nn_integral_eq_integral[OF i pos]  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2228  | 
by (subst SUP_eq_LIMSEQ) (auto simp: mono_def intro!: integral_mono_AE ilim)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2229  | 
finally have "(\<integral>\<^sup>+ x. ereal (u x) \<partial>M) = ereal x" .  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2230  | 
moreover have "(\<integral>\<^sup>+ x. ereal (- u x) \<partial>M) = 0"  | 
| 56996 | 2231  | 
proof (subst 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
 | 
2232  | 
show "(\<lambda>x. ereal (- u 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
 | 
2233  | 
using 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
 | 
2234  | 
from mono pos[of 0] lim show "AE x in M. ereal (- u x) \<le> 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2235  | 
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
 | 
2236  | 
fix x assume "mono (\<lambda>n. f n x)" "0 \<le> f 0 x" "(\<lambda>i. f i x) ----> u x"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2237  | 
then show "ereal (- u x) \<le> 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2238  | 
using incseq_le[of "\<lambda>n. f n x" "u x" 0] by (simp add: mono_def incseq_def)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2239  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2240  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2241  | 
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
 | 
2242  | 
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
 | 
2243  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2244  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2245  | 
lemma  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2246  | 
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
 | 
2247  | 
assumes f: "\<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
 | 
2248  | 
and lim: "AE x in M. (\<lambda>i. f i x) ----> u x"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2249  | 
and ilim: "(\<lambda>i. integral\<^sup>L M (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
 | 
2250  | 
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
 | 
2251  | 
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
 | 
2252  | 
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
 | 
2253  | 
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
 | 
2254  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2255  | 
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
 | 
2256  | 
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
 | 
2257  | 
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
 | 
2258  | 
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
 | 
2259  | 
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
 | 
2260  | 
using mono by (auto simp: field_simps 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
 | 
2261  | 
have 4: "AE x in M. (\<lambda>i. f i x - f 0 x) ----> u 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
 | 
2262  | 
using lim 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
 | 
2263  | 
have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x \<partial>M)) ----> x - integral\<^sup>L M (f 0)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2264  | 
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
 | 
2265  | 
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
 | 
2266  | 
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
 | 
2267  | 
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
 | 
2268  | 
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
 | 
2269  | 
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
 | 
2270  | 
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
 | 
2271  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2272  | 
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
 | 
2273  | 
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
 | 
2274  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2275  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2276  | 
lemma integral_norm_eq_0_iff:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2277  | 
  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
 | 
2278  | 
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
 | 
2279  | 
  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
 | 
2280  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2281  | 
have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>x. norm (f x) \<partial>M)"  | 
| 56996 | 2282  | 
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
 | 
2283  | 
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
 | 
2284  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2285  | 
  also have "\<dots> \<longleftrightarrow> emeasure M {x\<in>space M. ereal (norm (f x)) \<noteq> 0} = 0"
 | 
| 56996 | 2286  | 
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
 | 
2287  | 
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
 | 
2288  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2289  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2290  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2291  | 
lemma 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
 | 
2292  | 
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
 | 
2293  | 
  shows "integrable M f \<Longrightarrow> (\<integral>x. abs (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
 | 
2294  | 
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
 | 
2295  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2296  | 
lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\<lambda>x. a)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2297  | 
using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2298  | 
|
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2299  | 
lemma lebesgue_integral_const[simp]:  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2300  | 
  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
 | 
2301  | 
shows "(\<integral>x. a \<partial>M) = measure M (space M) *\<^sub>R a"  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2302  | 
proof -  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2303  | 
  { 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
 | 
2304  | 
then have ?thesis  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2305  | 
by (simp add: not_integrable_integral_eq measure_def integrable_iff_bounded) }  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2306  | 
moreover  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2307  | 
  { 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
 | 
2308  | 
moreover  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2309  | 
  { 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
 | 
2310  | 
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
 | 
2311  | 
proof qed fact  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2312  | 
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
 | 
2313  | 
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
 | 
2314  | 
also have "\<dots> = measure M (space M) *\<^sub>R a"  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2315  | 
by simp  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2316  | 
finally have ?thesis . }  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2317  | 
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
 | 
2318  | 
qed  | 
| 
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2319  | 
|
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2320  | 
lemma (in finite_measure) integrable_const_bound:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2321  | 
  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
 | 
2322  | 
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
 | 
2323  | 
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
 | 
2324  | 
apply assumption  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2325  | 
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
 | 
2326  | 
apply auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2327  | 
done  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2328  | 
|
| 59000 | 2329  | 
lemma integral_indicator_finite_real:  | 
2330  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
2331  | 
assumes [simp]: "finite A"  | 
|
2332  | 
  assumes [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
 | 
|
2333  | 
  assumes finite: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} < \<infinity>"
 | 
|
2334  | 
  shows "(\<integral>x. f x * indicator A x \<partial>M) = (\<Sum>a\<in>A. f a * measure M {a})"
 | 
|
2335  | 
proof -  | 
|
2336  | 
  have "(\<integral>x. f x * indicator A x \<partial>M) = (\<integral>x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
 | 
|
2337  | 
proof (intro integral_cong refl)  | 
|
2338  | 
    fix x show "f x * indicator A x = (\<Sum>a\<in>A. f a * indicator {a} x)"
 | 
|
2339  | 
by (auto split: split_indicator simp: eq_commute[of x] cong: conj_cong)  | 
|
2340  | 
qed  | 
|
2341  | 
  also have "\<dots> = (\<Sum>a\<in>A. f a * measure M {a})"
 | 
|
2342  | 
using finite by (subst integral_setsum) (auto simp add: integrable_mult_left_iff)  | 
|
2343  | 
finally show ?thesis .  | 
|
2344  | 
qed  | 
|
2345  | 
||
2346  | 
lemma (in finite_measure) ereal_integral_real:  | 
|
2347  | 
assumes [measurable]: "f \<in> borel_measurable M"  | 
|
2348  | 
assumes ae: "AE x in M. 0 \<le> f x" "AE x in M. f x \<le> ereal B"  | 
|
2349  | 
shows "ereal (\<integral>x. real (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"  | 
|
2350  | 
proof (subst nn_integral_eq_integral[symmetric])  | 
|
2351  | 
show "integrable M (\<lambda>x. real (f x))"  | 
|
2352  | 
using ae by (intro integrable_const_bound[where B=B]) (auto simp: real_le_ereal_iff)  | 
|
2353  | 
show "AE x in M. 0 \<le> real (f x)"  | 
|
2354  | 
using ae by (auto simp: real_of_ereal_pos)  | 
|
2355  | 
show "(\<integral>\<^sup>+ x. ereal (real (f x)) \<partial>M) = integral\<^sup>N M f"  | 
|
2356  | 
using ae by (intro nn_integral_cong_AE) (auto simp: ereal_real)  | 
|
2357  | 
qed  | 
|
2358  | 
||
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2359  | 
lemma (in finite_measure) integral_less_AE:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2360  | 
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
 | 
2361  | 
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
 | 
2362  | 
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
 | 
2363  | 
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
 | 
2364  | 
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
 | 
2365  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2366  | 
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
 | 
2367  | 
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
 | 
2368  | 
moreover  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2369  | 
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
 | 
2370  | 
proof  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2371  | 
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
 | 
2372  | 
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
 | 
2373  | 
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
 | 
2374  | 
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
 | 
2375  | 
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
 | 
2376  | 
    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
 | 
2377  | 
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
 | 
2378  | 
moreover  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2379  | 
    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 | 2380  | 
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
 | 
2381  | 
    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
 | 
2382  | 
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
 | 
2383  | 
ultimately have "emeasure M A = 0"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2384  | 
using emeasure_nonneg[of M A] by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2385  | 
with `(emeasure M) A \<noteq> 0` show False by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2386  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2387  | 
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
 | 
2388  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2389  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2390  | 
lemma (in finite_measure) integral_less_AE_space:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2391  | 
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
 | 
2392  | 
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
 | 
2393  | 
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
 | 
2394  | 
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
 | 
2395  | 
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
 | 
2396  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2397  | 
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
 | 
2398  | 
  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
 | 
| 59048 | 2399  | 
assumes [measurable_cong]: "sets M = sets borel" and f[measurable]: "integrable 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
 | 
2400  | 
  shows "((\<lambda>y. \<integral> x. indicator {.. y} x *\<^sub>R f x \<partial>M) ---> \<integral> x. f x \<partial>M) at_top"
 | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2401  | 
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
 | 
2402  | 
fix X :: "nat \<Rightarrow> real" assume "filterlim X at_top sequentially"  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2403  | 
  show "(\<lambda>n. \<integral>x. indicator {..X n} x *\<^sub>R f x \<partial>M) ----> integral\<^sup>L M f"
 | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2404  | 
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
 | 
2405  | 
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
 | 
2406  | 
by (rule integrable_norm) fact  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2407  | 
    show "AE x in M. (\<lambda>n. indicator {..X n} x *\<^sub>R f x) ----> f x"
 | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2408  | 
proof  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2409  | 
fix x  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2410  | 
from `filterlim X at_top sequentially`  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2411  | 
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
 | 
2412  | 
unfolding filterlim_at_top_ge[where c=x] by auto  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2413  | 
      then show "(\<lambda>n. indicator {..X n} x *\<^sub>R f x) ----> f x"
 | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2414  | 
by (intro Lim_eventually) (auto split: split_indicator elim!: eventually_elim1)  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2415  | 
qed  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
2416  | 
    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
 | 
2417  | 
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
 | 
2418  | 
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
 | 
2419  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2420  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2421  | 
lemma  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2422  | 
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
 | 
2423  | 
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
 | 
2424  | 
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
 | 
2425  | 
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
 | 
2426  | 
  assumes int: "\<And>y. integrable M (\<lambda>x. f x * indicator {.. y} x)"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2427  | 
  assumes conv: "((\<lambda>y. \<integral> x. f x * indicator {.. y} x \<partial>M) ---> x) at_top"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2428  | 
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
 | 
2429  | 
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
 | 
2430  | 
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
 | 
2431  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2432  | 
  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
 | 
2433  | 
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
 | 
2434  | 
  { fix x have "eventually (\<lambda>n. f x * indicator {..real n} x = f x) sequentially"
 | 
| 
59587
 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 
nipkow 
parents: 
59452 
diff
changeset
 | 
2435  | 
by (rule eventually_sequentiallyI[of "nat(ceiling x)"])  | 
| 
 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 
nipkow 
parents: 
59452 
diff
changeset
 | 
2436  | 
(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
 | 
2437  | 
from filterlim_cong[OF refl refl this]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2438  | 
  have "AE x in M. (\<lambda>i. f x * indicator {..real i} x) ----> f x"
 | 
| 
58729
 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 
hoelzl 
parents: 
57514 
diff
changeset
 | 
2439  | 
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
 | 
2440  | 
  have "(\<lambda>i. \<integral> x. f x * indicator {..real i} x \<partial>M) ----> x"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2441  | 
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
 | 
2442  | 
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
 | 
2443  | 
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
 | 
2444  | 
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
 | 
2445  | 
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
 | 
2446  | 
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
 | 
2447  | 
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
 | 
2448  | 
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
 | 
2449  | 
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
 | 
2450  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2451  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2452  | 
subsection {* Product measure *}
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2453  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2454  | 
lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2455  | 
  fixes f :: "_ \<Rightarrow> _ \<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
 | 
2456  | 
assumes [measurable]: "split f \<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
 | 
2457  | 
shows "Measurable.pred N (\<lambda>x. 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
 | 
2458  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2459  | 
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
 | 
2460  | 
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
 | 
2461  | 
show ?thesis  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2462  | 
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
 | 
2463  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2464  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2465  | 
lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2466  | 
"(\<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
 | 
2467  | 
    {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
 | 
2468  | 
(\<lambda>x. measure M (A x)) \<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  | 
unfolding measure_def by (intro measurable_emeasure borel_measurable_real_of_ereal)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2470  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2471  | 
lemma Collect_subset [simp]: "{x\<in>A. P x} \<subseteq> 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
 | 
2472  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2473  | 
lemma (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2474  | 
  fixes f :: "_ \<Rightarrow> _ \<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
 | 
2475  | 
assumes f[measurable]: "split f \<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
 | 
2476  | 
shows "(\<lambda>x. \<integral>y. f x y \<partial>M) \<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
 | 
2477  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2478  | 
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
 | 
2479  | 
then have s: "\<And>i. simple_function (N \<Otimes>\<^sub>M 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
 | 
2480  | 
"\<And>x y. x \<in> space N \<Longrightarrow> y \<in> space M \<Longrightarrow> (\<lambda>i. s i (x, 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
 | 
2481  | 
"\<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)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2482  | 
by (auto simp: space_pair_measure norm_conv_dist)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2483  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2484  | 
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
 | 
2485  | 
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
 | 
2486  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2487  | 
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
 | 
2488  | 
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
 | 
2489  | 
|
| 
57166
 
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
2490  | 
def f' \<equiv> "\<lambda>i x. if integrable M (f x) then simple_bochner_integral M (\<lambda>y. s i (x, y)) 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
 | 
2491  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2492  | 
  { 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
 | 
2493  | 
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
 | 
2494  | 
      (\<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
 | 
2495  | 
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
 | 
2496  | 
unfolding simple_bochner_integral_def  | 
| 57418 | 2497  | 
by (intro setsum.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
 | 
2498  | 
(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
 | 
2499  | 
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
 | 
2500  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2501  | 
show ?thesis  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2502  | 
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
 | 
2503  | 
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
 | 
2504  | 
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
 | 
2505  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2506  | 
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
 | 
2507  | 
    { 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
 | 
2508  | 
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
 | 
2509  | 
by (intro integrable_norm integrable_mult_right int_f)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2510  | 
have "(\<lambda>i. integral\<^sup>L M (\<lambda>y. s i (x, y))) ----> 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
 | 
2511  | 
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
 | 
2512  | 
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
 | 
2513  | 
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
 | 
2514  | 
using x by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2515  | 
show "AE xa in M. (\<lambda>i. s i (x, xa)) ----> 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
 | 
2516  | 
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
 | 
2517  | 
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
 | 
2518  | 
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
 | 
2519  | 
qed fact  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2520  | 
moreover  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2521  | 
      { fix i
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2522  | 
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
 | 
2523  | 
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
 | 
2524  | 
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
 | 
2525  | 
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
 | 
2526  | 
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
 | 
2527  | 
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
 | 
2528  | 
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
 | 
2529  | 
(auto simp: space_pair_measure dest: finite_subset)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2530  | 
have "(\<integral>\<^sup>+ y. ereal (norm (s i (x, y))) \<partial>M) \<le> (\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M)"  | 
| 56996 | 2531  | 
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
 | 
2532  | 
also have "(\<integral>\<^sup>+ y. 2 * 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
 | 
2533  | 
using int_2f by (simp add: 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
 | 
2534  | 
finally show "(\<integral>\<^sup>+ xa. ereal (norm (s i (x, xa))) \<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
 | 
2535  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2536  | 
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
 | 
2537  | 
by (rule simple_bochner_integrable_eq_integral[symmetric]) }  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2538  | 
ultimately have "(\<lambda>i. simple_bochner_integral M (\<lambda>y. s i (x, y))) ----> 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
 | 
2539  | 
by simp }  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2540  | 
then  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2541  | 
show "(\<lambda>i. f' i x) ----> 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
 | 
2542  | 
unfolding f'_def  | 
| 
58729
 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 
hoelzl 
parents: 
57514 
diff
changeset
 | 
2543  | 
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
 | 
2544  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2545  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2546  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2547  | 
lemma (in pair_sigma_finite) integrable_product_swap:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2548  | 
  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
 | 
2549  | 
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
 | 
2550  | 
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
 | 
2551  | 
proof -  | 
| 61169 | 2552  | 
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
 | 
2553  | 
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
 | 
2554  | 
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
 | 
2555  | 
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
 | 
2556  | 
(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
 | 
2557  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2558  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2559  | 
lemma (in pair_sigma_finite) integrable_product_swap_iff:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2560  | 
  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
 | 
2561  | 
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
 | 
2562  | 
proof -  | 
| 61169 | 2563  | 
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
 | 
2564  | 
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
 | 
2565  | 
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
 | 
2566  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2567  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2568  | 
lemma (in pair_sigma_finite) integral_product_swap:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2569  | 
  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
 | 
2570  | 
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
 | 
2571  | 
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
 | 
2572  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2573  | 
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
 | 
2574  | 
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
 | 
2575  | 
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
 | 
2576  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2577  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2578  | 
lemma (in pair_sigma_finite) emeasure_pair_measure_finite:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2579  | 
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
 | 
2580  | 
  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
 | 
2581  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2582  | 
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
 | 
2583  | 
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
 | 
2584  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2585  | 
then have "AE x in M1. emeasure M2 (Pair x -` A) \<noteq> \<infinity>"  | 
| 56996 | 2586  | 
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
 | 
2587  | 
  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
 | 
2588  | 
using sets.sets_into_space[OF A] 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
 | 
2589  | 
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
 | 
2590  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2591  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2592  | 
lemma (in pair_sigma_finite) AE_integrable_fst':  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2593  | 
  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
 | 
2594  | 
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
 | 
2595  | 
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
 | 
2596  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2597  | 
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 | 2598  | 
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
 | 
2599  | 
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
 | 
2600  | 
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
 | 
2601  | 
finally have "AE x in M1. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<noteq> \<infinity>"  | 
| 56996 | 2602  | 
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
 | 
2603  | 
(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
 | 
2604  | 
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
 | 
2605  | 
by eventually_elim  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2606  | 
(auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]])  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2607  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2608  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2609  | 
lemma (in pair_sigma_finite) integrable_fst':  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2610  | 
  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
 | 
2611  | 
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
 | 
2612  | 
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
 | 
2613  | 
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
 | 
2614  | 
proof  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2615  | 
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
 | 
2616  | 
by (rule M2.borel_measurable_lebesgue_integral) simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2617  | 
have "(\<integral>\<^sup>+ x. ereal (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)"  | 
| 56996 | 2618  | 
using AE_integrable_fst'[OF f] by (auto intro!: nn_integral_mono_AE integral_norm_bound_ereal)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2619  | 
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 | 2620  | 
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
 | 
2621  | 
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
 | 
2622  | 
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
 | 
2623  | 
finally show "(\<integral>\<^sup>+ x. ereal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) < \<infinity>" .  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2624  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2625  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2626  | 
lemma (in pair_sigma_finite) integral_fst':  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2627  | 
  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
 | 
2628  | 
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
 | 
2629  | 
shows "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>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
 | 
2630  | 
using f proof induct  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2631  | 
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
 | 
2632  | 
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
 | 
2633  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2634  | 
  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
 | 
2635  | 
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
 | 
2636  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2637  | 
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
 | 
2638  | 
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
 | 
2639  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2640  | 
  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
 | 
2641  | 
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
 | 
2642  | 
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
 | 
2643  | 
    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
 | 
2644  | 
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
 | 
2645  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2646  | 
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
 | 
2647  | 
proof (subst 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
 | 
2648  | 
    have "(\<integral>\<^sup>+x. ereal (measure 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
 | 
2649  | 
      (\<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
 | 
2650  | 
using emeasure_pair_measure_finite[OF base]  | 
| 56996 | 2651  | 
by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ereal_measure)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2652  | 
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
 | 
2653  | 
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
 | 
2654  | 
by (subst M2.emeasure_pair_measure_alt)  | 
| 56996 | 2655  | 
(auto intro!: nn_integral_cong arg_cong[where f="emeasure M2"] 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
 | 
2656  | 
    finally have *: "(\<integral>\<^sup>+x. ereal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) = 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
 | 
2657  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2658  | 
    from base * show "integrable M1 (\<lambda>x. measure M2 {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
 | 
2659  | 
by (simp add: measure_nonneg 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
 | 
2660  | 
    then have "(\<integral>x. measure 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
 | 
2661  | 
      (\<integral>\<^sup>+x. ereal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1)"
 | 
| 56996 | 2662  | 
by (rule nn_integral_eq_integral[symmetric]) (simp add: measure_nonneg)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2663  | 
also note *  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2664  | 
    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"
 | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2665  | 
using base by (simp add: emeasure_eq_ereal_measure)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2666  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2667  | 
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
 | 
2668  | 
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
 | 
2669  | 
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
 | 
2670  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2671  | 
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
 | 
2672  | 
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
 | 
2673  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2674  | 
have "(\<integral> x. \<integral> y. f (x, 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
 | 
2675  | 
(\<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
 | 
2676  | 
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
 | 
2677  | 
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
 | 
2678  | 
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
 | 
2679  | 
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
 | 
2680  | 
apply simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2681  | 
done  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2682  | 
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
 | 
2683  | 
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
 | 
2684  | 
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
 | 
2685  | 
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
 | 
2686  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2687  | 
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
 | 
2688  | 
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
 | 
2689  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2690  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2691  | 
show ?case  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2692  | 
proof (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
 | 
2693  | 
show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) ----> 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
 | 
2694  | 
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
 | 
2695  | 
show "integrable (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. 2 * norm (f x))"  | 
| 57036 | 2696  | 
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
 | 
2697  | 
qed (insert lim, auto)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2698  | 
have "(\<lambda>i. \<integral> x. \<integral> y. s i (x, y) \<partial>M2 \<partial>M1) ----> \<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
 | 
2699  | 
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
 | 
2700  | 
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
 | 
2701  | 
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
 | 
2702  | 
with AE_space AE_integrable_fst'[OF lim(5)]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2703  | 
show "AE x in M1. (\<lambda>i. \<integral> y. s i (x, y) \<partial>M2) ----> \<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
 | 
2704  | 
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
 | 
2705  | 
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
 | 
2706  | 
s: "\<forall>i. 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
 | 
2707  | 
show "(\<lambda>i. \<integral> y. s i (x, y) \<partial>M2) ----> \<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
 | 
2708  | 
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
 | 
2709  | 
show "integrable M2 (\<lambda>y. 2 * norm (f (x, y)))"  | 
| 57036 | 2710  | 
using f 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
 | 
2711  | 
show "AE xa in M2. (\<lambda>i. s i (x, xa)) ----> 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
 | 
2712  | 
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
 | 
2713  | 
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
 | 
2714  | 
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
 | 
2715  | 
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
 | 
2716  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2717  | 
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
 | 
2718  | 
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
 | 
2719  | 
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
 | 
2720  | 
using AE_space AE_integrable_fst'[OF lim(1), of i] AE_integrable_fst'[OF lim(5)]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2721  | 
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
 | 
2722  | 
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
 | 
2723  | 
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
 | 
2724  | 
from s have "norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral>\<^sup>+y. norm (s i (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
 | 
2725  | 
by (rule integral_norm_bound_ereal)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2726  | 
also have "\<dots> \<le> (\<integral>\<^sup>+y. 2 * norm (f (x, y)) \<partial>M2)"  | 
| 56996 | 2727  | 
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
 | 
2728  | 
also have "\<dots> = (\<integral>y. 2 * norm (f (x, y)) \<partial>M2)"  | 
| 56996 | 2729  | 
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
 | 
2730  | 
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
 | 
2731  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2732  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2733  | 
qed simp_all  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2734  | 
then show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) ----> \<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
 | 
2735  | 
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
 | 
2736  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2737  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2738  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2739  | 
lemma (in pair_sigma_finite)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2740  | 
  fixes f :: "_ \<Rightarrow> _ \<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
 | 
2741  | 
assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (split f)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2742  | 
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
 | 
2743  | 
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
 | 
2744  | 
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
 | 
2745  | 
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
 | 
2746  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2747  | 
lemma (in pair_sigma_finite)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2748  | 
  fixes f :: "_ \<Rightarrow> _ \<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
 | 
2749  | 
assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) (split f)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2750  | 
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
 | 
2751  | 
and integrable_snd: "integrable M2 (\<lambda>y. \<integral>x. f x y \<partial>M1)" (is "?INT")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2752  | 
and integral_snd: "(\<integral>y. (\<integral>x. f x y \<partial>M1) \<partial>M2) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (split f)" (is "?EQ")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2753  | 
proof -  | 
| 61169 | 2754  | 
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
 | 
2755  | 
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
 | 
2756  | 
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
 | 
2757  | 
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
 | 
2758  | 
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
 | 
2759  | 
show ?EQ using Q.integral_fst'[OF Q_int]  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2760  | 
using integral_product_swap[of "split 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
 | 
2761  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2762  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2763  | 
lemma (in pair_sigma_finite) Fubini_integral:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2764  | 
  fixes f :: "_ \<Rightarrow> _ \<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
 | 
2765  | 
assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (split f)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2766  | 
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
 | 
2767  | 
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
 | 
2768  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2769  | 
lemma (in product_sigma_finite) product_integral_singleton:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2770  | 
  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
 | 
2771  | 
  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
 | 
2772  | 
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
 | 
2773  | 
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
 | 
2774  | 
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
 | 
2775  | 
done  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2776  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2777  | 
lemma (in product_sigma_finite) product_integral_fold:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2778  | 
  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
 | 
2779  | 
  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
 | 
2780  | 
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
 | 
2781  | 
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
 | 
2782  | 
proof -  | 
| 61169 | 2783  | 
interpret I: finite_product_sigma_finite M I by standard fact  | 
2784  | 
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
 | 
2785  | 
have "finite (I \<union> J)" using fin by auto  | 
| 61169 | 2786  | 
interpret IJ: finite_product_sigma_finite M "I \<union> J" by standard fact  | 
2787  | 
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
 | 
2788  | 
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
 | 
2789  | 
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
 | 
2790  | 
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
 | 
2791  | 
by auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2792  | 
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
 | 
2793  | 
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
 | 
2794  | 
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
 | 
2795  | 
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
 | 
2796  | 
show ?thesis  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2797  | 
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
 | 
2798  | 
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
 | 
2799  | 
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
 | 
2800  | 
apply simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2801  | 
done  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2802  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2803  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2804  | 
lemma (in product_sigma_finite) product_integral_insert:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2805  | 
  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
 | 
2806  | 
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
 | 
2807  | 
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
 | 
2808  | 
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)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2809  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2810  | 
  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
 | 
2811  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2812  | 
  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
 | 
2813  | 
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
 | 
2814  | 
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
 | 
2815  | 
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
 | 
2816  | 
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
 | 
2817  | 
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
 | 
2818  | 
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
 | 
2819  | 
show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2820  | 
using measurable_comp[OF measurable_component_update f_borel, OF x `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
 | 
2821  | 
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
 | 
2822  | 
    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
 | 
2823  | 
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
 | 
2824  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2825  | 
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
 | 
2826  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2827  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2828  | 
lemma (in product_sigma_finite) product_integrable_setprod:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2829  | 
  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
 | 
2830  | 
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
 | 
2831  | 
shows "integrable (Pi\<^sub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f")  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2832  | 
proof (unfold integrable_iff_bounded, intro conjI)  | 
| 61169 | 2833  | 
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
 | 
2834  | 
|
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2835  | 
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
 | 
2836  | 
using assms by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2837  | 
have "(\<integral>\<^sup>+ x. ereal (norm (\<Prod>i\<in>I. f i (x 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
 | 
2838  | 
(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. ereal (norm (f i (x 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
 | 
2839  | 
by (simp add: setprod_norm setprod_ereal)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2840  | 
also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+ x. ereal (norm (f i x)) \<partial>M i)"  | 
| 56996 | 2841  | 
using assms by (intro product_nn_integral_setprod) auto  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2842  | 
also have "\<dots> < \<infinity>"  | 
| 56996 | 2843  | 
using integrable by (simp add: setprod_PInf nn_integral_nonneg 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
 | 
2844  | 
finally show "(\<integral>\<^sup>+ x. ereal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) < \<infinity>" .  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2845  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2846  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2847  | 
lemma (in product_sigma_finite) product_integral_setprod:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2848  | 
  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
 | 
2849  | 
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
 | 
2850  | 
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))"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2851  | 
using assms proof induct  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2852  | 
case empty  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2853  | 
  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
 | 
2854  | 
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
 | 
2855  | 
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
 | 
2856  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2857  | 
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
 | 
2858  | 
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
 | 
2859  | 
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
 | 
2860  | 
integrable (Pi\<^sub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2861  | 
by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset)  | 
| 61169 | 2862  | 
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
 | 
2863  | 
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))"  | 
| 57418 | 2864  | 
using `i \<notin> I` by (auto intro!: setprod.cong)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2865  | 
show ?case  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2866  | 
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
 | 
2867  | 
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
 | 
2868  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2869  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2870  | 
lemma integrable_subalgebra:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2871  | 
  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
 | 
2872  | 
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
 | 
2873  | 
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
 | 
2874  | 
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
 | 
2875  | 
proof -  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2876  | 
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
 | 
2877  | 
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
 | 
2878  | 
with assms show ?thesis  | 
| 56996 | 2879  | 
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
 | 
2880  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2881  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2882  | 
lemma integral_subalgebra:  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2883  | 
  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
 | 
2884  | 
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
 | 
2885  | 
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
 | 
2886  | 
shows "integral\<^sup>L N 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
 | 
2887  | 
proof cases  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2888  | 
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
 | 
2889  | 
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
 | 
2890  | 
proof induct  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2891  | 
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
 | 
2892  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2893  | 
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
 | 
2894  | 
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
 | 
2895  | 
by simp  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2896  | 
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
 | 
2897  | 
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
 | 
2898  | 
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
 | 
2899  | 
next  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2900  | 
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
 | 
2901  | 
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
 | 
2902  | 
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
 | 
2903  | 
show ?case  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2904  | 
proof (intro LIMSEQ_unique)  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2905  | 
show "(\<lambda>i. integral\<^sup>L N (s i)) ----> integral\<^sup>L N f"  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2906  | 
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
 | 
2907  | 
using lim  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2908  | 
apply auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2909  | 
done  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2910  | 
show "(\<lambda>i. integral\<^sup>L N (s i)) ----> 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
 | 
2911  | 
unfolding lim  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2912  | 
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
 | 
2913  | 
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
 | 
2914  | 
apply auto  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2915  | 
done  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2916  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2917  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2918  | 
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
 | 
2919  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2920  | 
hide_const simple_bochner_integral  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2921  | 
hide_const 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
 | 
2922  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents:  
diff
changeset
 | 
2923  | 
end  | 
| 
59353
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
2924  |