author | eberlm <eberlm@in.tum.de> |
Fri, 11 Aug 2017 14:29:30 +0200 | |
changeset 66394 | 32084d7e6b59 |
parent 65680 | 378a2f11bec9 |
child 66447 | a1f5c5c26fa6 |
permissions | -rw-r--r-- |
63627 | 1 |
(* Title: HOL/Analysis/Bochner_Integration.thy |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2 |
Author: Johannes Hölzl, TU München |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3 |
*) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
4 |
|
61808 | 5 |
section \<open>Bochner Integration for Vector-Valued Functions\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
6 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
7 |
theory Bochner_Integration |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
8 |
imports Finite_Product_Measure |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
9 |
begin |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
10 |
|
61808 | 11 |
text \<open> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
12 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
13 |
In the following development of the Bochner integral we use second countable topologies instead |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
14 |
of separable spaces. A second countable topology is also separable. |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
15 |
|
61808 | 16 |
\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
17 |
|
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" |
61969 | 21 |
shows "\<exists>F. (\<forall>i. simple_function M (F i)) \<and> (\<forall>x\<in>space M. (\<lambda>i. F i x) \<longlonglongrightarrow> f x) \<and> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
22 |
(\<forall>i. \<forall>x\<in>space M. dist (F i x) z \<le> 2 * dist (f x) z)" |
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 |
|
63040 | 27 |
define e where "e = from_nat_into D" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
28 |
{ fix n x |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
29 |
obtain d where "d \<in> D" and d: "d \<in> ball x (1 / Suc n)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
30 |
using D[of "ball x (1 / Suc n)"] by auto |
61808 | 31 |
from \<open>d \<in> D\<close> D[of UNIV] \<open>countable D\<close> obtain i where "d = e i" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
32 |
unfolding e_def by (auto dest: from_nat_into_surj) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
33 |
with d have "\<exists>i. dist x (e i) < 1 / Suc n" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
34 |
by auto } |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
35 |
note e = this |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
36 |
|
63040 | 37 |
define A where [abs_def]: "A m n = |
38 |
{x\<in>space M. dist (f x) (e n) < 1 / (Suc m) \<and> 1 / (Suc m) \<le> dist (f x) z}" for m n |
|
39 |
define B where [abs_def]: "B m = disjointed (A m)" for m |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
40 |
|
63040 | 41 |
define m where [abs_def]: "m N x = Max {m. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}" for N x |
42 |
define F where [abs_def]: "F N x = |
|
43 |
(if (\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)) \<and> (\<exists>n\<le>N. x \<in> B (m N x) n) |
|
44 |
then e (LEAST n. x \<in> B (m N x) n) else z)" for N x |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
45 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
46 |
have B_imp_A[intro, simp]: "\<And>x m n. x \<in> B m n \<Longrightarrow> x \<in> A m n" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
47 |
using disjointed_subset[of "A m" for m] unfolding B_def by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
48 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
49 |
{ fix m |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
50 |
have "\<And>n. A m n \<in> sets M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
51 |
by (auto simp: A_def) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
52 |
then have "\<And>n. B m n \<in> sets M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
53 |
using sets.range_disjointed_sets[of "A m" M] by (auto simp: B_def) } |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
54 |
note this[measurable] |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
55 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
56 |
{ fix N i x assume "\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
57 |
then have "m N x \<in> {m::nat. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
58 |
unfolding m_def by (intro Max_in) auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
59 |
then have "m N x \<le> N" "\<exists>n\<le>N. x \<in> B (m N x) n" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
60 |
by auto } |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
61 |
note m = this |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
62 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
63 |
{ fix j N i x assume "j \<le> N" "i \<le> N" "x \<in> B j i" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
64 |
then have "j \<le> m N x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
65 |
unfolding m_def by (intro Max_ge) auto } |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
66 |
note m_upper = this |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
67 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
68 |
show ?thesis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
69 |
unfolding simple_function_def |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
70 |
proof (safe intro!: exI[of _ F]) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
71 |
have [measurable]: "\<And>i. F i \<in> borel_measurable M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
72 |
unfolding F_def m_def by measurable |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
73 |
show "\<And>x i. F i -` {x} \<inter> space M \<in> sets M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
74 |
by measurable |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
75 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
76 |
{ fix i |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
77 |
{ fix n x assume "x \<in> B (m i x) n" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
78 |
then have "(LEAST n. x \<in> B (m i x) n) \<le> n" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
79 |
by (intro Least_le) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
80 |
also assume "n \<le> i" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
81 |
finally have "(LEAST n. x \<in> B (m i x) n) \<le> i" . } |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
82 |
then have "F i ` space M \<subseteq> {z} \<union> e ` {.. i}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
83 |
by (auto simp: F_def) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
84 |
then show "finite (F i ` space M)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
85 |
by (rule finite_subset) auto } |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
86 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
87 |
{ fix N i n x assume "i \<le> N" "n \<le> N" "x \<in> B i n" |
60585 | 88 |
then have 1: "\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)" by auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
89 |
from m[OF this] obtain n where n: "m N x \<le> N" "n \<le> N" "x \<in> B (m N x) n" by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
90 |
moreover |
63040 | 91 |
define L where "L = (LEAST n. x \<in> B (m N x) n)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
92 |
have "dist (f x) (e L) < 1 / Suc (m N x)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
93 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
94 |
have "x \<in> B (m N x) L" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
95 |
using n(3) unfolding L_def by (rule LeastI) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
96 |
then have "x \<in> A (m N x) L" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
97 |
by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
98 |
then show ?thesis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
99 |
unfolding A_def by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
100 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
101 |
ultimately have "dist (f x) (F N x) < 1 / Suc (m N x)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
102 |
by (auto simp add: F_def L_def) } |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
103 |
note * = this |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
104 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
105 |
fix x assume "x \<in> space M" |
61969 | 106 |
show "(\<lambda>i. F i x) \<longlonglongrightarrow> f x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
107 |
proof cases |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
108 |
assume "f x = z" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
109 |
then have "\<And>i n. x \<notin> A i n" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
110 |
unfolding A_def by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
111 |
then have "\<And>i. F i x = z" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
112 |
by (auto simp: F_def) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
113 |
then show ?thesis |
61808 | 114 |
using \<open>f x = z\<close> by auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
115 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
116 |
assume "f x \<noteq> z" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
117 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
118 |
show ?thesis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
119 |
proof (rule tendstoI) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
120 |
fix e :: real assume "0 < e" |
61808 | 121 |
with \<open>f x \<noteq> z\<close> obtain n where "1 / Suc n < e" "1 / Suc n < dist (f x) z" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
122 |
by (metis dist_nz order_less_trans neq_iff nat_approx_posE) |
61808 | 123 |
with \<open>x\<in>space M\<close> \<open>f x \<noteq> z\<close> have "x \<in> (\<Union>i. B n i)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
124 |
unfolding A_def B_def UN_disjointed_eq using e by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
125 |
then obtain i where i: "x \<in> B n i" by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
126 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
127 |
show "eventually (\<lambda>i. dist (F i x) (f x) < e) sequentially" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
128 |
using eventually_ge_at_top[of "max n i"] |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
129 |
proof eventually_elim |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
130 |
fix j assume j: "max n i \<le> j" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
131 |
with i have "dist (f x) (F j x) < 1 / Suc (m j x)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
132 |
by (intro *[OF _ _ i]) auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
133 |
also have "\<dots> \<le> 1 / Suc n" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
134 |
using j m_upper[OF _ _ i] |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
135 |
by (auto simp: field_simps) |
61808 | 136 |
also note \<open>1 / Suc n < e\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
137 |
finally show "dist (F j x) (f x) < e" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
138 |
by (simp add: less_imp_le dist_commute) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
139 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
140 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
141 |
qed |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
142 |
fix i |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
143 |
{ fix n m assume "x \<in> A n m" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
144 |
then have "dist (e m) (f x) + dist (f x) z \<le> 2 * dist (f x) z" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
145 |
unfolding A_def by (auto simp: dist_commute) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
146 |
also have "dist (e m) z \<le> dist (e m) (f x) + dist (f x) z" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
147 |
by (rule dist_triangle) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
148 |
finally (xtrans) have "dist (e m) z \<le> 2 * dist (f x) z" . } |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
149 |
then show "dist (F i x) z \<le> 2 * dist (f x) z" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
150 |
unfolding F_def |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
151 |
apply auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
152 |
apply (rule LeastI2) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
153 |
apply auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
154 |
done |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
155 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
156 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
157 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
158 |
lemma |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
159 |
fixes f :: "'a \<Rightarrow> 'b::semiring_1" assumes "finite A" |
64267 | 160 |
shows sum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator (B x) (g x)) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)" |
161 |
and sum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator (B x) (g x) * f x) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
162 |
unfolding indicator_def |
64267 | 163 |
using assms by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
164 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
165 |
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
|
166 |
fixes P :: "('a \<Rightarrow> real) \<Rightarrow> bool" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
167 |
assumes u: "u \<in> borel_measurable M" "\<And>x. 0 \<le> u x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
168 |
assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
169 |
assumes mult: "\<And>u c. 0 \<le> c \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
170 |
assumes add: "\<And>u v. u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> v x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)" |
61969 | 171 |
assumes seq: "\<And>U. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. 0 \<le> U i x) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. U i x) \<longlonglongrightarrow> u x) \<Longrightarrow> P u" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
172 |
shows "P u" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
173 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
174 |
have "(\<lambda>x. ennreal (u x)) \<in> borel_measurable M" using u by auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
175 |
from borel_measurable_implies_simple_function_sequence'[OF this] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
176 |
obtain U where U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i x. U i x < top" and |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
177 |
sup: "\<And>x. (SUP i. U i x) = ennreal (u x)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
178 |
by blast |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
179 |
|
63040 | 180 |
define U' where [abs_def]: "U' i x = indicator (space M) x * enn2real (U i x)" for i x |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
181 |
then have U'_sf[measurable]: "\<And>i. simple_function M (U' i)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
182 |
using U by (auto intro!: simple_function_compose1[where g=enn2real]) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
183 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
184 |
show "P u" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
185 |
proof (rule seq) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
186 |
show U': "U' i \<in> borel_measurable M" "\<And>x. 0 \<le> U' i x" for i |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
187 |
using U by (auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
188 |
intro: borel_measurable_simple_function |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
189 |
intro!: borel_measurable_enn2real borel_measurable_times |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63627
diff
changeset
|
190 |
simp: U'_def zero_le_mult_iff) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
191 |
show "incseq U'" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
192 |
using U(2,3) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
193 |
by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def enn2real_mono) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
194 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
195 |
fix x assume x: "x \<in> space M" |
61969 | 196 |
have "(\<lambda>i. U i x) \<longlonglongrightarrow> (SUP i. U i x)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
197 |
using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_def le_fun_def) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
198 |
moreover have "(\<lambda>i. U i x) = (\<lambda>i. ennreal (U' i x))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
199 |
using x U(3) by (auto simp: fun_eq_iff U'_def image_iff eq_commute) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
200 |
moreover have "(SUP i. U i x) = ennreal (u x)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
201 |
using sup u(2) by (simp add: max_def) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
202 |
ultimately show "(\<lambda>i. U' i x) \<longlonglongrightarrow> u x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
203 |
using u U' by simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
204 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
205 |
fix i |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
206 |
have "U' i ` space M \<subseteq> enn2real ` (U i ` space M)" "finite (U i ` space M)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
207 |
unfolding U'_def using U(1) by (auto dest: simple_functionD) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
208 |
then have fin: "finite (U' i ` space M)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
209 |
by (metis finite_subset finite_imageI) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
210 |
moreover have "\<And>z. {y. U' i z = y \<and> y \<in> U' i ` space M \<and> z \<in> space M} = (if z \<in> space M then {U' i z} else {})" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
211 |
by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
212 |
ultimately have U': "(\<lambda>z. \<Sum>y\<in>U' i`space M. y * indicator {x\<in>space M. U' i x = y} z) = U' i" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
213 |
by (simp add: U'_def fun_eq_iff) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
214 |
have "\<And>x. x \<in> U' i ` space M \<Longrightarrow> 0 \<le> x" |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63627
diff
changeset
|
215 |
by (auto simp: U'_def) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
216 |
with fin have "P (\<lambda>z. \<Sum>y\<in>U' i`space M. y * indicator {x\<in>space M. U' i x = y} z)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
217 |
proof induct |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
218 |
case empty from set[of "{}"] show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
219 |
by (simp add: indicator_def[abs_def]) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
220 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
221 |
case (insert x F) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
222 |
then show ?case |
64267 | 223 |
by (auto intro!: add mult set sum_nonneg split: split_indicator split_indicator_asm |
224 |
simp del: sum_mult_indicator simp: sum_nonneg_eq_0_iff) |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
225 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
226 |
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
|
227 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
228 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
229 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
230 |
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
|
231 |
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
|
232 |
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
|
233 |
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
|
234 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
235 |
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
|
236 |
"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
|
237 |
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
|
238 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
239 |
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
|
240 |
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
|
241 |
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
|
242 |
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
|
243 |
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
|
244 |
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
|
245 |
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
|
246 |
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
|
247 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
248 |
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
|
249 |
"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
|
250 |
"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
|
251 |
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
|
252 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
253 |
assume fin: "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" "emeasure M {y \<in> space M. g y \<noteq> 0} \<noteq> \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
254 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
255 |
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
|
256 |
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
|
257 |
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
|
258 |
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
|
259 |
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
|
260 |
finally show "emeasure M {y \<in> space M. p (f y) (g y) \<noteq> 0} \<noteq> \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
261 |
using fin by (auto simp: top_unique) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
262 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
263 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
264 |
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
|
265 |
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
|
266 |
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
|
267 |
proof cases |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
268 |
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
|
269 |
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
|
270 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
271 |
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
|
272 |
|
63040 | 273 |
define m where "m = Min (f`space M - {0})" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
274 |
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
|
275 |
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
|
276 |
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
|
277 |
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
|
278 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
279 |
from m have "m * emeasure M {x\<in>space M. 0 \<noteq> f x} = |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
280 |
(\<integral>\<^sup>+x. m * indicator {x\<in>space M. 0 \<noteq> f x} x \<partial>M)" |
56996 | 281 |
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
|
282 |
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
|
283 |
using AE_space |
56996 | 284 |
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
|
285 |
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
|
286 |
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
|
287 |
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
|
288 |
qed |
61808 | 289 |
also note \<open>\<dots> < \<infinity>\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
290 |
finally show ?thesis |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
291 |
using m by (auto simp: ennreal_mult_less_top) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
292 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
293 |
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
|
294 |
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
|
295 |
by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
296 |
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
|
297 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
298 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
299 |
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
|
300 |
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
|
301 |
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
|
302 |
proof |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
303 |
have "emeasure M {y \<in> space M. ennreal (norm (f y)) \<noteq> 0} \<noteq> \<infinity>" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
304 |
proof (rule simple_function_finite_support) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
305 |
show "simple_function M (\<lambda>x. ennreal (norm (f x)))" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
306 |
using f by (rule simple_function_compose1) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
307 |
show "(\<integral>\<^sup>+ y. ennreal (norm (f y)) \<partial>M) < \<infinity>" by fact |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
308 |
qed simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
309 |
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
|
310 |
qed fact |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
311 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
312 |
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
|
313 |
"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
|
314 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
315 |
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
|
316 |
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
|
317 |
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
|
318 |
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
|
319 |
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
|
320 |
(is "_ = ?r") |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
321 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
322 |
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
|
323 |
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
|
324 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
325 |
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
|
326 |
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
|
327 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
328 |
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
|
329 |
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
|
330 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
331 |
{ 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
|
332 |
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
|
333 |
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
|
334 |
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
|
335 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
336 |
have "simple_bochner_integral M f = |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
337 |
(\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
338 |
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
|
339 |
unfolding simple_bochner_integral_def |
64267 | 340 |
proof (safe intro!: sum.cong scaleR_cong_right) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
341 |
fix y assume y: "y \<in> space M" "f y \<noteq> 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
342 |
have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} = |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
343 |
{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
|
344 |
by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
345 |
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
|
346 |
(\<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
|
347 |
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
|
348 |
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
|
349 |
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
|
350 |
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
|
351 |
moreover |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
352 |
{ 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
|
353 |
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
|
354 |
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
|
355 |
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
|
356 |
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
|
357 |
then have "emeasure M {xa \<in> space M. g xa = g x} < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
358 |
using f by (auto simp: simple_bochner_integrable.simps less_top) } |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
359 |
ultimately |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
360 |
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
|
361 |
(\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then measure M {x \<in> space M. g x = z} else 0)" |
64267 | 362 |
apply (simp add: sum.If_cases eq) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
363 |
apply (subst measure_finite_Union[symmetric]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
364 |
apply (auto simp: disjoint_family_on_def less_top) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
365 |
done |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
366 |
qed |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
367 |
also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
368 |
if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} *\<^sub>R y else 0))" |
64267 | 369 |
by (auto intro!: sum.cong simp: scaleR_sum_left) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
370 |
also have "\<dots> = ?r" |
64267 | 371 |
by (subst sum.commute) |
372 |
(auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq) |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
373 |
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
|
374 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
375 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
376 |
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
|
377 |
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
|
378 |
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
|
379 |
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
|
380 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
381 |
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
|
382 |
(\<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
|
383 |
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
|
384 |
(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
|
385 |
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
|
386 |
(\<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
|
387 |
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
|
388 |
(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
|
389 |
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
|
390 |
(\<Sum>y\<in>(\<lambda>x. (f x, g x)) ` space M. measure M {x \<in> space M. (f x, g x) = y} *\<^sub>R snd y)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
391 |
by (intro simple_bochner_integral_partition) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
392 |
(auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
393 |
ultimately show ?thesis |
64267 | 394 |
by (simp add: sum.distrib[symmetric] scaleR_add_right) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
395 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
396 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
397 |
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
|
398 |
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
|
399 |
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
|
400 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
401 |
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
|
402 |
(\<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
|
403 |
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
|
404 |
(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
|
405 |
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
|
406 |
also have "\<dots> = f (simple_bochner_integral M g)" |
64267 | 407 |
by (simp add: simple_bochner_integral_def sum scaleR) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
408 |
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
|
409 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
410 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
411 |
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
|
412 |
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
|
413 |
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
|
414 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
415 |
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
|
416 |
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
|
417 |
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
|
418 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
419 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
420 |
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
|
421 |
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
|
422 |
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
|
423 |
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
|
424 |
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
|
425 |
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
|
426 |
(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
|
427 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
428 |
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
|
429 |
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
|
430 |
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
|
431 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
432 |
have "norm (simple_bochner_integral M f) \<le> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
433 |
(\<Sum>y\<in>f ` space M. norm (measure M {x \<in> space M. f x = y} *\<^sub>R y))" |
64267 | 434 |
unfolding simple_bochner_integral_def by (rule norm_sum) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
435 |
also have "\<dots> = (\<Sum>y\<in>f ` space M. measure M {x \<in> space M. f x = y} *\<^sub>R norm y)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
436 |
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
|
437 |
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
|
438 |
using f |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
439 |
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
|
440 |
(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
|
441 |
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
|
442 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
443 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
444 |
lemma simple_bochner_integral_nonneg[simp]: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
445 |
fixes f :: "'a \<Rightarrow> real" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
446 |
shows "(\<And>x. 0 \<le> f x) \<Longrightarrow> 0 \<le> simple_bochner_integral M f" |
65680
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
paulson <lp15@cam.ac.uk>
parents:
64911
diff
changeset
|
447 |
by (force simp add: simple_bochner_integral_def intro: sum_nonneg) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
448 |
|
56996 | 449 |
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
|
450 |
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
|
451 |
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
|
452 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
453 |
{ fix x y z have "(x \<noteq> 0 \<Longrightarrow> y = z) \<Longrightarrow> ennreal x * y = ennreal x * z" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
454 |
by (cases "x = 0") (auto simp: zero_ennreal_def[symmetric]) } |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
455 |
note ennreal_cong_mult = this |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
456 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
457 |
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
|
458 |
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
|
459 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
460 |
{ fix y assume y: "y \<in> space M" "f y \<noteq> 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
461 |
have "ennreal (measure M {x \<in> space M. f x = f y}) = emeasure M {x \<in> space M. f x = f y}" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
462 |
proof (rule emeasure_eq_ennreal_measure[symmetric]) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
463 |
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
|
464 |
using y by (intro emeasure_mono) auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
465 |
with f show "emeasure M {x \<in> space M. f x = f y} \<noteq> top" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
466 |
by (auto simp: simple_bochner_integrable.simps top_unique) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
467 |
qed |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
468 |
moreover have "{x \<in> space M. f x = f y} = (\<lambda>x. ennreal (f x)) -` {ennreal (f y)} \<inter> space M" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
469 |
using f by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
470 |
ultimately have "ennreal (measure M {x \<in> space M. f x = f y}) = |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
471 |
emeasure M ((\<lambda>x. ennreal (f x)) -` {ennreal (f y)} \<inter> space M)" by simp } |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
472 |
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
|
473 |
unfolding simple_integral_def |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
474 |
by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ennreal (f x)" and v=enn2real]) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
475 |
(auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases |
64267 | 476 |
intro!: sum.cong ennreal_cong_mult |
477 |
simp: sum_ennreal[symmetric] ac_simps ennreal_mult |
|
478 |
simp del: sum_ennreal) |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
479 |
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
|
480 |
using f |
56996 | 481 |
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
|
482 |
(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
|
483 |
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
|
484 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
485 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
486 |
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
|
487 |
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
|
488 |
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
|
489 |
assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
490 |
shows "ennreal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \<le> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
491 |
(\<integral>\<^sup>+ x. norm (f x - s x) \<partial>M) + (\<integral>\<^sup>+ x. norm (f x - t x) \<partial>M)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
492 |
(is "ennreal (norm (?s - ?t)) \<le> ?S + ?T") |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
493 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
494 |
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
|
495 |
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
|
496 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
497 |
have "ennreal (norm (?s - ?t)) = norm (simple_bochner_integral M (\<lambda>x. s x - t x))" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
498 |
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
|
499 |
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
|
500 |
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
|
501 |
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
|
502 |
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
|
503 |
using simple_bochner_integrable_compose2[of "\<lambda>x y. norm (x - y)" M "s" "t"] s t |
56996 | 504 |
by (auto intro!: simple_bochner_integral_eq_nn_integral) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
505 |
also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
506 |
by (auto intro!: nn_integral_mono simp: ennreal_plus[symmetric] simp del: ennreal_plus) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
507 |
(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
|
508 |
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
|
509 |
also have "\<dots> = ?S + ?T" |
56996 | 510 |
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
|
511 |
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
|
512 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
513 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
514 |
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
|
515 |
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
|
516 |
"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
|
517 |
(\<And>i. simple_bochner_integrable M (s i)) \<Longrightarrow> |
61969 | 518 |
(\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0 \<Longrightarrow> |
519 |
(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x \<Longrightarrow> |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
520 |
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
|
521 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
522 |
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
|
523 |
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
|
524 |
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
|
525 |
unfolding has_bochner_integral.simps assms(1,3) |
56996 | 526 |
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
|
527 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
528 |
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
|
529 |
"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
|
530 |
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
|
531 |
unfolding has_bochner_integral.simps |
61969 | 532 |
by (intro arg_cong[where f=Ex] ext conj_cong rev_conj_cong refl arg_cong[where f="\<lambda>x. x \<longlonglongrightarrow> 0"] |
56996 | 533 |
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
|
534 |
auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
535 |
|
59353
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59048
diff
changeset
|
536 |
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
|
537 |
"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
|
538 |
by (rule has_bochner_integral.cases) |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59048
diff
changeset
|
539 |
|
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59048
diff
changeset
|
540 |
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
|
541 |
"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
|
542 |
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
|
543 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
544 |
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
|
545 |
"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
|
546 |
by (rule has_bochner_integral.intros[where s="\<lambda>_. f"]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
547 |
(auto intro: borel_measurable_simple_function |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
548 |
elim: simple_bochner_integrable.cases |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
549 |
simp: zero_ennreal_def[symmetric]) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
550 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
551 |
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
|
552 |
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
|
553 |
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
|
554 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
555 |
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
|
556 |
proof |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
557 |
have "{y \<in> space M. (indicator A y::real) \<noteq> 0} = A" |
61808 | 558 |
using sets.sets_into_space[OF \<open>A\<in>sets M\<close>] by (auto split: split_indicator) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
559 |
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
|
560 |
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
|
561 |
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
|
562 |
moreover have "simple_bochner_integral M (indicator A) = measure M A" |
56996 | 563 |
using simple_bochner_integral_eq_nn_integral[OF sbi] A |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
564 |
by (simp add: ennreal_indicator emeasure_eq_ennreal_measure) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
565 |
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
|
566 |
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
|
567 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
568 |
|
57036 | 569 |
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
|
570 |
"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
|
571 |
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
|
572 |
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
|
573 |
fix sf sg |
61969 | 574 |
assume f_sf: "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - sf i x) \<partial>M) \<longlonglongrightarrow> 0" |
575 |
assume g_sg: "(\<lambda>i. \<integral>\<^sup>+ x. norm (g x - sg i x) \<partial>M) \<longlonglongrightarrow> 0" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
576 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
577 |
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
|
578 |
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
|
579 |
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
|
580 |
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
|
581 |
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
|
582 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
583 |
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
|
584 |
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
|
585 |
|
61969 | 586 |
show "(\<lambda>i. \<integral>\<^sup>+ x. (norm (f x + g x - (sf i x + sg i x))) \<partial>M) \<longlonglongrightarrow> 0" |
587 |
(is "?f \<longlonglongrightarrow> 0") |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
588 |
proof (rule tendsto_sandwich) |
61969 | 589 |
show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
590 |
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
|
591 |
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
|
592 |
(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
|
593 |
proof (intro always_eventually allI) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
594 |
fix i have "?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
595 |
by (auto intro!: nn_integral_mono norm_diff_triangle_ineq |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
596 |
simp del: ennreal_plus simp add: ennreal_plus[symmetric]) |
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 |
61969 | 601 |
show "?g \<longlonglongrightarrow> 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
602 |
using tendsto_add[OF f_sf g_sg] by simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
|
61969 | 617 |
fix s assume f_s: "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
|
61969 | 628 |
show "(\<lambda>i. \<integral>\<^sup>+ x. norm (T (f x) - T (s i x)) \<partial>M) \<longlonglongrightarrow> 0" |
629 |
(is "?f \<longlonglongrightarrow> 0") |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
630 |
proof (rule tendsto_sandwich) |
61969 | 631 |
show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
632 |
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
|
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) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
637 |
fix i have "?f i \<le> (\<integral>\<^sup>+ x. ennreal K * norm (f x - s i x) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
638 |
using K by (intro nn_integral_mono) (auto simp: ac_simps ennreal_mult[symmetric]) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
61969 | 643 |
show "?g \<longlonglongrightarrow> 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
644 |
using ennreal_tendsto_cmult[OF _ f_s] by simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
|
61969 | 647 |
assume "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x" |
648 |
with s show "(\<lambda>i. simple_bochner_integral M (\<lambda>x. T (s i x))) \<longlonglongrightarrow> T x" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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"] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
654 |
simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
675 |
lemmas has_bochner_integral_divide = |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
|
64267 | 717 |
lemma has_bochner_integral_sum: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
727 |
lim_0: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
728 |
from order_tendstoD[OF lim_0, of "\<infinity>"] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
729 |
obtain i where f_s_fin: "(\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) < \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
730 |
by (auto simp: eventually_sequentially) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
731 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
732 |
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
|
733 |
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
|
734 |
|
63040 | 735 |
define m where "m = (if space M = {} then 0 else Max ((\<lambda>x. norm (s i x))`space M))" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
736 |
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
|
737 |
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
|
738 |
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
|
739 |
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
|
740 |
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
|
741 |
by (auto simp: m_def image_comp comp_def Max_ge_iff) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
742 |
then have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal m * indicator {x\<in>space M. s i x \<noteq> 0} x \<partial>M)" |
56996 | 743 |
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
|
744 |
also have "\<dots> < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
745 |
using s by (subst nn_integral_cmult_indicator) (auto simp: \<open>0 \<le> m\<close> simple_bochner_integrable.simps ennreal_mult_less_top less_top) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
746 |
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
|
747 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
748 |
have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
749 |
by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
750 |
(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 |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
755 |
finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" . |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
61969 | 763 |
x: "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x" (is "?s \<longlonglongrightarrow> x") and |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
764 |
s[simp]: "\<And>i. simple_bochner_integrable M (s i)" and |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
765 |
lim: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" and |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
773 |
show "(\<lambda>i. ennreal (norm (?s i))) \<longlonglongrightarrow> norm x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
774 |
using x by (auto simp: tendsto_ennreal_iff intro: tendsto_intros) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
779 |
have "ennreal (norm (?s n)) \<le> simple_bochner_integral M (\<lambda>x. norm (s n x))" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
784 |
also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
785 |
by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric]) |
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) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
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 |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
791 |
have "?t \<longlonglongrightarrow> 0 + (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
792 |
using has_bochner_integral_implies_finite_norm[OF i] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
793 |
by (intro tendsto_add tendsto_const lim) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
794 |
then show "?t \<longlonglongrightarrow> \<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
61969 | 805 |
assume "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?S \<longlonglongrightarrow> 0") |
806 |
assume "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - t i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?T \<longlonglongrightarrow> 0") |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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)" |
61969 | 815 |
assume "?s \<longlonglongrightarrow> x" "?t \<longlonglongrightarrow> y" |
816 |
then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> norm (x - y)" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
819 |
have "(\<lambda>i. ennreal (norm (?s i - ?t i))) \<longlonglongrightarrow> ennreal 0" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
820 |
proof (rule tendsto_sandwich) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
821 |
show "eventually (\<lambda>i. 0 \<le> ennreal (norm (?s i - ?t i))) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> ennreal 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
822 |
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
|
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) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
826 |
show "(\<lambda>i. ?S i + ?T i) \<longlonglongrightarrow> ennreal 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
827 |
using tendsto_add[OF \<open>?S \<longlonglongrightarrow> 0\<close> \<open>?T \<longlonglongrightarrow> 0\<close>] by simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
828 |
qed |
61969 | 829 |
then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
830 |
by (simp add: ennreal_0[symmetric] del: ennreal_0) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
831 |
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
|
832 |
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
|
833 |
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
|
834 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
835 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
836 |
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
|
837 |
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
|
838 |
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
|
839 |
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
|
840 |
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
|
841 |
using f |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
842 |
proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
843 |
fix s assume "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
844 |
also have "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) = (\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
845 |
using ae |
56996 | 846 |
by (intro ext nn_integral_cong_AE, eventually_elim) simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
847 |
finally show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" . |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
848 |
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
|
849 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
850 |
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
|
851 |
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
|
852 |
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
|
853 |
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
|
854 |
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
|
855 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
856 |
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
|
857 |
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
|
858 |
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
|
859 |
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
|
860 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
861 |
|
57137 | 862 |
lemma simple_bochner_integrable_restrict_space: |
863 |
fixes f :: "_ \<Rightarrow> 'b::real_normed_vector" |
|
864 |
assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" |
|
865 |
shows "simple_bochner_integrable (restrict_space M \<Omega>) f \<longleftrightarrow> |
|
866 |
simple_bochner_integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" |
|
867 |
by (simp add: simple_bochner_integrable.simps space_restrict_space |
|
868 |
simple_function_restrict_space[OF \<Omega>] emeasure_restrict_space[OF \<Omega>] Collect_restrict |
|
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63627
diff
changeset
|
869 |
indicator_eq_0_iff conj_left_commute) |
57137 | 870 |
|
871 |
lemma simple_bochner_integral_restrict_space: |
|
872 |
fixes f :: "_ \<Rightarrow> 'b::real_normed_vector" |
|
873 |
assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" |
|
874 |
assumes f: "simple_bochner_integrable (restrict_space M \<Omega>) f" |
|
875 |
shows "simple_bochner_integral (restrict_space M \<Omega>) f = |
|
876 |
simple_bochner_integral M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" |
|
877 |
proof - |
|
878 |
have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x)`space M)" |
|
879 |
using f simple_bochner_integrable_restrict_space[OF \<Omega>, of f] |
|
880 |
by (simp add: simple_bochner_integrable.simps simple_function_def) |
|
881 |
then show ?thesis |
|
882 |
by (auto simp: space_restrict_space measure_restrict_space[OF \<Omega>(1)] le_infI2 |
|
883 |
simple_bochner_integral_def Collect_restrict |
|
884 |
split: split_indicator split_indicator_asm |
|
64267 | 885 |
intro!: sum.mono_neutral_cong_left arg_cong2[where f=measure]) |
57137 | 886 |
qed |
887 |
||
61681
ca53150406c9
option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents:
61609
diff
changeset
|
888 |
context |
62093 | 889 |
notes [[inductive_internals]] |
61681
ca53150406c9
option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents:
61609
diff
changeset
|
890 |
begin |
ca53150406c9
option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents:
61609
diff
changeset
|
891 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
892 |
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
|
893 |
"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
|
894 |
|
61681
ca53150406c9
option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents:
61609
diff
changeset
|
895 |
end |
ca53150406c9
option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents:
61609
diff
changeset
|
896 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
897 |
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
|
898 |
"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
|
899 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
900 |
syntax |
59357
f366643536cd
allow line breaks in integral notation
Andreas Lochbihler
parents:
59353
diff
changeset
|
901 |
"_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
|
902 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
903 |
translations |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
904 |
"\<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
|
905 |
|
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
906 |
syntax |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
907 |
"_ascii_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real" ("(3LINT (1_)/|(_)./ _)" [0,110,60] 60) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
908 |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
909 |
translations |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
910 |
"LINT x|M. f" == "CONST lebesgue_integral M (\<lambda>x. f)" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
911 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
912 |
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
|
913 |
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
|
914 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
915 |
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
|
916 |
"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
|
917 |
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
|
918 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
919 |
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
|
920 |
"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
|
921 |
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
|
922 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
923 |
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
|
924 |
"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
|
925 |
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
|
926 |
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
|
927 |
|
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
928 |
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
|
929 |
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
|
930 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
931 |
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
|
932 |
"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
|
933 |
(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
|
934 |
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
|
935 |
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
|
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 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
|
938 |
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
|
939 |
|
59353
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59048
diff
changeset
|
940 |
lemma borel_measurable_integrable'[measurable_dest]: |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59048
diff
changeset
|
941 |
"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
|
942 |
using borel_measurable_integrable[measurable] by measurable |
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents:
59048
diff
changeset
|
943 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
944 |
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
|
945 |
"M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> integrable N g" |
63092 | 946 |
by (simp cong: has_bochner_integral_cong add: integrable.simps) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
947 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
948 |
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
|
949 |
"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
|
950 |
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
|
951 |
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
|
952 |
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
|
953 |
|
64008
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
954 |
lemma integrable_cong_AE_imp: |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
955 |
"integrable M g \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> (AE x in M. g x = f x) \<Longrightarrow> integrable M f" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
956 |
using integrable_cong_AE[of f M g] by (auto simp: eq_commute) |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
957 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
958 |
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
|
959 |
"M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integral\<^sup>L M f = integral\<^sup>L N g" |
63566 | 960 |
by (simp cong: has_bochner_integral_cong cong del: if_weak_cong add: lebesgue_integral_def) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
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 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
|
963 |
"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
|
964 |
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
|
965 |
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
|
966 |
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
|
967 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
968 |
lemma integrable_add[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x + g x)" |
57036 | 969 |
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
|
970 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
971 |
lemma integrable_zero[simp, intro]: "integrable M (\<lambda>x. 0)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
972 |
by (metis has_bochner_integral_zero integrable.simps) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
973 |
|
64267 | 974 |
lemma integrable_sum[simp, intro]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> integrable M (\<lambda>x. \<Sum>i\<in>I. f i x)" |
975 |
by (metis has_bochner_integral_sum integrable.simps) |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
976 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
977 |
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
|
978 |
integrable M (\<lambda>x. indicator A x *\<^sub>R c)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
979 |
by (metis has_bochner_integral_indicator integrable.simps) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
980 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
981 |
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
|
982 |
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
|
983 |
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
|
984 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
985 |
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
|
986 |
by (auto simp: integrable.simps intro: has_bochner_integral_diff) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
987 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
988 |
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
|
989 |
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
|
990 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
991 |
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
|
992 |
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
|
993 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
994 |
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
|
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_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
|
998 |
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
|
999 |
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
|
1000 |
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
|
1001 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1002 |
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
|
1003 |
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
|
1004 |
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
|
1005 |
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
|
1006 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1007 |
lemma integrable_divide_zero[simp, intro]: |
59867
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
haftmann
parents:
59587
diff
changeset
|
1008 |
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
|
1009 |
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
|
1010 |
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
|
1011 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1012 |
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
|
1013 |
"(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
|
1014 |
unfolding integrable.simps by fastforce |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1015 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1016 |
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
|
1017 |
"(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
|
1018 |
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
|
1019 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1020 |
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
|
1021 |
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
|
1022 |
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
|
1023 |
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
|
1024 |
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
|
1025 |
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
|
1026 |
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
|
1027 |
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
|
1028 |
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
|
1029 |
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
|
1030 |
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
|
1031 |
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
|
1032 |
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
|
1033 |
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
|
1034 |
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
|
1035 |
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
|
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_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
|
1038 |
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
|
1039 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1040 |
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
|
1041 |
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
|
1042 |
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
|
1043 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1044 |
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
|
1045 |
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
|
1046 |
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
|
1047 |
|
64267 | 1048 |
lemma integral_sum: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1049 |
integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))" |
64267 | 1050 |
by (intro has_bochner_integral_integral_eq has_bochner_integral_sum has_bochner_integral_integrable) |
1051 |
||
1052 |
lemma integral_sum'[simp]: "(\<And>i. i \<in> I =simp=> integrable M (f i)) \<Longrightarrow> |
|
62083 | 1053 |
integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))" |
64267 | 1054 |
unfolding simp_implies_def by (rule integral_sum) |
62083 | 1055 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1056 |
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
|
1057 |
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
|
1058 |
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
|
1059 |
|
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1060 |
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
|
1061 |
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
|
1062 |
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
|
1063 |
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
|
1064 |
proof cases |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1065 |
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
|
1066 |
by simp |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1067 |
next |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1068 |
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
|
1069 |
show ?thesis |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1070 |
proof cases |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1071 |
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
|
1072 |
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
|
1073 |
next |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1074 |
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
|
1075 |
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
|
1076 |
proof |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1077 |
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
|
1078 |
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
|
1079 |
show False |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1080 |
by auto |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1081 |
qed |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1082 |
ultimately show ?thesis |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1083 |
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
|
1084 |
qed |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1085 |
qed |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1086 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1087 |
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
|
1088 |
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
|
1089 |
|
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1090 |
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
|
1091 |
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
|
1092 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1093 |
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
|
1094 |
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
|
1095 |
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
|
1096 |
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
|
1097 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1098 |
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
|
1099 |
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
|
1100 |
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
|
1101 |
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
|
1102 |
|
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1103 |
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
|
1104 |
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
|
1105 |
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
|
1106 |
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
|
1107 |
|
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1108 |
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
|
1109 |
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
|
1110 |
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
|
1111 |
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
|
1112 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1113 |
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
|
1114 |
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
|
1115 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1116 |
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
|
1117 |
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
|
1118 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1119 |
lemma integral_divide_zero[simp]: |
59867
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
haftmann
parents:
59587
diff
changeset
|
1120 |
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
|
1121 |
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
|
1122 |
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
|
1123 |
|
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1124 |
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
|
1125 |
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
|
1126 |
|
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1127 |
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
|
1128 |
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
|
1129 |
|
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1130 |
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
|
1131 |
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
|
1132 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1133 |
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
|
1134 |
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
|
1135 |
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
|
1136 |
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
|
1137 |
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
|
1138 |
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
|
1139 |
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
|
1140 |
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
|
1141 |
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
|
1142 |
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
|
1143 |
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
|
1144 |
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
|
1145 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1146 |
lemma integral_norm_bound_ennreal: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1147 |
"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
|
1148 |
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
|
1149 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1150 |
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
|
1151 |
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
|
1152 |
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
|
1153 |
assumes s: "\<And>i. simple_bochner_integrable M (s i)" |
61969 | 1154 |
assumes lim: "(\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?S \<longlonglongrightarrow> 0") |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1155 |
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
|
1156 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1157 |
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
|
1158 |
|
61969 | 1159 |
have "\<exists>x. ?s \<longlonglongrightarrow> x" |
64287 | 1160 |
unfolding convergent_eq_Cauchy |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1161 |
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
|
1162 |
fix e :: real assume "0 < e" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1163 |
then have "0 < ennreal (e / 2)" by auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1164 |
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
|
1165 |
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
|
1166 |
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
|
1167 |
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
|
1168 |
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
|
1169 |
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
|
1170 |
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
|
1171 |
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
|
1172 |
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
|
1173 |
by (intro simple_bochner_integral_bounded s f) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1174 |
also have "\<dots> < ennreal (e / 2) + e / 2" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1175 |
by (intro add_strict_mono M n m) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1176 |
also have "\<dots> = e" using \<open>0<e\<close> by (simp del: ennreal_plus add: ennreal_plus[symmetric]) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1177 |
finally show "dist (?s n) (?s m) < e" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1178 |
using \<open>0<e\<close> by (simp add: dist_norm ennreal_less_iff) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1179 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1180 |
qed |
61969 | 1181 |
then obtain x where "?s \<longlonglongrightarrow> x" .. |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1182 |
show ?thesis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1183 |
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
|
1184 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1185 |
|
56996 | 1186 |
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
|
1187 |
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
|
1188 |
assumes [measurable]: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1189 |
"\<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
|
1190 |
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
|
1191 |
and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>" |
61969 | 1192 |
and u': "AE x in M. (\<lambda>i. u i x) \<longlonglongrightarrow> u' x" |
1193 |
shows "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) \<longlonglongrightarrow> 0" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1194 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1195 |
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
|
1196 |
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
|
1197 |
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
|
1198 |
proof (eventually_elim, intro allI) |
61969 | 1199 |
fix i x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x" "\<forall>j. norm (u j x) \<le> w x" "\<forall>j. norm (u j x) \<le> w x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1200 |
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
|
1201 |
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
|
1202 |
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
|
1203 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1204 |
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
|
1205 |
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
|
1206 |
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
|
1207 |
qed |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1208 |
have w_nonneg: "AE x in M. 0 \<le> w x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1209 |
using bound[of 0] by (auto intro: order_trans[OF norm_ge_zero]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1210 |
|
61969 | 1211 |
have "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. 0 \<partial>M)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1212 |
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
|
1213 |
show "(\<integral>\<^sup>+x. 2 * w x \<partial>M) < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1214 |
by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) (insert w_nonneg, auto simp: ennreal_mult ) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1215 |
show "AE x in M. (\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1216 |
using u' |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1217 |
proof eventually_elim |
61969 | 1218 |
fix x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1219 |
from tendsto_diff[OF tendsto_const[of "u' x"] this] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1220 |
show "(\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1221 |
by (simp add: tendsto_norm_zero_iff ennreal_0[symmetric] del: ennreal_0) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1222 |
qed |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1223 |
qed (insert bnd w_nonneg, auto) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1224 |
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
|
1225 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1226 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1227 |
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
|
1228 |
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
|
1229 |
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
|
1230 |
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
|
1231 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1232 |
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
|
1233 |
s: "\<And>i. simple_function M (s i)" and |
61969 | 1234 |
pointwise: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" and |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1235 |
bound: "\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)" |
62379
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents:
62093
diff
changeset
|
1236 |
by simp metis |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1237 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1238 |
show ?thesis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1239 |
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
|
1240 |
{ fix i |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1241 |
have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)" |
56996 | 1242 |
by (intro nn_integral_mono) (simp add: bound) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1243 |
also have "\<dots> = 2 * (\<integral>\<^sup>+x. ennreal (norm (f x)) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1244 |
by (simp add: ennreal_mult nn_integral_cmult) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1245 |
also have "\<dots> < top" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1246 |
using fin by (simp add: ennreal_mult_less_top) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1247 |
finally have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1248 |
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
|
1249 |
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
|
1250 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1251 |
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
|
1252 |
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
|
1253 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1254 |
show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" |
56996 | 1255 |
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
|
1256 |
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
|
1257 |
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
|
1258 |
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
|
1259 |
using s by (auto intro: borel_measurable_simple_function) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1260 |
show "(\<integral>\<^sup>+ x. ennreal (2 * norm (f x)) \<partial>M) < \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1261 |
using fin by (simp add: nn_integral_cmult ennreal_mult ennreal_mult_less_top) |
61969 | 1262 |
show "AE x in M. (\<lambda>i. s i x) \<longlonglongrightarrow> f x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1263 |
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
|
1264 |
qed fact |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1265 |
qed fact |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1266 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1267 |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1268 |
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
|
1269 |
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
|
1270 |
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
|
1271 |
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
|
1272 |
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
|
1273 |
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
|
1274 |
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
|
1275 |
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
|
1276 |
{ 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
|
1277 |
using norm_ge_zero[of x] by arith } |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1278 |
with bnd null have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (max 0 B) * indicator A x \<partial>M)" |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1279 |
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
|
1280 |
also have "\<dots> < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1281 |
using finite by (subst nn_integral_cmult_indicator) (auto simp: ennreal_mult_less_top) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1282 |
finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" . |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1283 |
qed simp |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1284 |
|
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1285 |
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
|
1286 |
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
|
1287 |
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
|
1288 |
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
|
1289 |
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
|
1290 |
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
|
1291 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1292 |
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
|
1293 |
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
|
1294 |
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
|
1295 |
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
|
1296 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1297 |
have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)" |
56996 | 1298 |
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
|
1299 |
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
|
1300 |
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
|
1301 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1302 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1303 |
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
|
1304 |
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
|
1305 |
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
|
1306 |
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
|
1307 |
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
|
1308 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1309 |
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
|
1310 |
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
|
1311 |
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
|
1312 |
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
|
1313 |
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
|
1314 |
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
|
1315 |
proof safe |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1316 |
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
|
1317 |
assume "AE x in M. norm (g x) \<le> norm (f x)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1318 |
then have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" |
56996 | 1319 |
by (intro nn_integral_mono_AE) auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1320 |
also assume "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1321 |
finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>" . |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1322 |
qed |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1323 |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1324 |
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
|
1325 |
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
|
1326 |
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
|
1327 |
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
|
1328 |
|
59000 | 1329 |
lemma integrable_real_mult_indicator: |
1330 |
fixes f :: "'a \<Rightarrow> real" |
|
1331 |
shows "A \<in> sets M \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. f x * indicator A x)" |
|
1332 |
using integrable_mult_indicator[of A M f] by (simp add: mult_ac) |
|
1333 |
||
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1334 |
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
|
1335 |
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
|
1336 |
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
|
1337 |
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
|
1338 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1339 |
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
|
1340 |
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1341 |
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
|
1342 |
using assms by (rule integrable_bound) auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1343 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1344 |
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
|
1345 |
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
|
1346 |
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
|
1347 |
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
|
1348 |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1349 |
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
|
1350 |
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
|
1351 |
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
|
1352 |
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
|
1353 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1354 |
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
|
1355 |
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
|
1356 |
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
|
1357 |
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
|
1358 |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1359 |
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
|
1360 |
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
|
1361 |
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
|
1362 |
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
|
1363 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1364 |
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
|
1365 |
fixes f :: "'a \<Rightarrow> real" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1366 |
assumes 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
|
1367 |
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
|
1368 |
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
|
1369 |
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
|
1370 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1371 |
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
|
1372 |
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
|
1373 |
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
|
1374 |
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
|
1375 |
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
|
1376 |
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
|
1377 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1378 |
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
|
1379 |
"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
|
1380 |
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
|
1381 |
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
|
1382 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1383 |
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
|
1384 |
"integrable M (indicator A::_ \<Rightarrow> real) \<longleftrightarrow> A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1385 |
by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator nn_integral_indicator' |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1386 |
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
|
1387 |
|
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1388 |
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
|
1389 |
proof cases |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1390 |
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
|
1391 |
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
|
1392 |
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
|
1393 |
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
|
1394 |
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
|
1395 |
finally show ?thesis . |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1396 |
next |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1397 |
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
|
1398 |
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
|
1399 |
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
|
1400 |
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
|
1401 |
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
|
1402 |
also have "\<dots> = measure M (A \<inter> space M)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1403 |
using * by (auto simp: measure_def emeasure_notin_sets not_less top_unique) |
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1404 |
finally show ?thesis . |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1405 |
qed |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
1406 |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1407 |
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
|
1408 |
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
|
1409 |
assumes X: "countable X" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1410 |
assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1411 |
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
|
1412 |
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
|
1413 |
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
|
1414 |
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
|
1415 |
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
|
1416 |
{ 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
|
1417 |
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
|
1418 |
moreover |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1419 |
{ 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
|
1420 |
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
|
1421 |
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
|
1422 |
next |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1423 |
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
|
1424 |
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
|
1425 |
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
|
1426 |
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
|
1427 |
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
|
1428 |
by simp |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1429 |
qed |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1430 |
|
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1431 |
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
|
1432 |
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
|
1433 |
assumes X: "countable X" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1434 |
assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1435 |
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
|
1436 |
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
|
1437 |
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
|
1438 |
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
|
1439 |
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
|
1440 |
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
|
1441 |
|
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1442 |
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
|
1443 |
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
|
1444 |
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
|
1445 |
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
|
1446 |
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
|
1447 |
|
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1448 |
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
|
1449 |
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
|
1450 |
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
|
1451 |
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
|
1452 |
qed |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1453 |
qed |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1454 |
|
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1455 |
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
|
1456 |
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
|
1457 |
assumes X: "countable X" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1458 |
assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1459 |
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
|
1460 |
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
|
1461 |
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
|
1462 |
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
|
1463 |
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
|
1464 |
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
|
1465 |
|
57137 | 1466 |
lemma |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1467 |
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
|
1468 |
assumes "f \<in> borel_measurable M" "\<And>i. s i \<in> borel_measurable M" "integrable M w" |
61969 | 1469 |
assumes lim: "AE x in M. (\<lambda>i. s i x) \<longlonglongrightarrow> f x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1470 |
assumes bound: "\<And>i. AE x in M. norm (s i x) \<le> w x" |
57137 | 1471 |
shows integrable_dominated_convergence: "integrable M f" |
1472 |
and integrable_dominated_convergence2: "\<And>i. integrable M (s i)" |
|
61969 | 1473 |
and integral_dominated_convergence: "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1474 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1475 |
have w_nonneg: "AE x in M. 0 \<le> w x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1476 |
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
|
1477 |
then have "(\<integral>\<^sup>+x. w x \<partial>M) = (\<integral>\<^sup>+x. norm (w x) \<partial>M)" |
56996 | 1478 |
by (intro nn_integral_cong_AE) auto |
61808 | 1479 |
with \<open>integrable M w\<close> have w: "w \<in> borel_measurable M" "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1480 |
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
|
1481 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1482 |
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
|
1483 |
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
|
1484 |
proof |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1485 |
fix i |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1486 |
have "(\<integral>\<^sup>+ x. ennreal (norm (s i x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1487 |
using bound[of i] w_nonneg by (intro nn_integral_mono_AE) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1488 |
with w show "(\<integral>\<^sup>+ x. ennreal (norm (s i x)) \<partial>M) < \<infinity>" by auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1489 |
qed fact |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1490 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1491 |
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
|
1492 |
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
|
1493 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1494 |
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
|
1495 |
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
|
1496 |
proof |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1497 |
have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1498 |
using all_bound lim w_nonneg |
56996 | 1499 |
proof (intro nn_integral_mono_AE, eventually_elim) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1500 |
fix x assume "\<forall>i. norm (s i x) \<le> w x" "(\<lambda>i. s i x) \<longlonglongrightarrow> f x" "0 \<le> w x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1501 |
then show "ennreal (norm (f x)) \<le> ennreal (w x)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1502 |
by (intro LIMSEQ_le_const2[where X="\<lambda>i. ennreal (norm (s i x))"]) (auto intro: tendsto_intros) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1503 |
qed |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1504 |
with w show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" by auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1505 |
qed fact |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1506 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1507 |
have "(\<lambda>n. ennreal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) \<longlonglongrightarrow> ennreal 0" (is "?d \<longlonglongrightarrow> ennreal 0") |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1508 |
proof (rule tendsto_sandwich) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1509 |
show "eventually (\<lambda>n. ennreal 0 \<le> ?d n) sequentially" "(\<lambda>_. ennreal 0) \<longlonglongrightarrow> ennreal 0" by auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1510 |
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
|
1511 |
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
|
1512 |
fix n |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1513 |
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
|
1514 |
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
|
1515 |
also have "\<dots> \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1516 |
by (intro int_f int_s integrable_diff integral_norm_bound_ennreal) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1517 |
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
|
1518 |
qed |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1519 |
show "(\<lambda>n. \<integral>\<^sup>+x. norm (s n x - f x) \<partial>M) \<longlonglongrightarrow> ennreal 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1520 |
unfolding ennreal_0 |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1521 |
apply (subst norm_minus_commute) |
56996 | 1522 |
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
|
1523 |
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
|
1524 |
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
|
1525 |
qed fact+ |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1526 |
qed |
61969 | 1527 |
then have "(\<lambda>n. integral\<^sup>L M (s n) - integral\<^sup>L M f) \<longlonglongrightarrow> 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1528 |
by (simp add: tendsto_norm_zero_iff del: ennreal_0) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1529 |
from tendsto_add[OF this tendsto_const[of "integral\<^sup>L M f"]] |
61969 | 1530 |
show "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f" by simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1531 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1532 |
|
61897
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1533 |
context |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1534 |
fixes s :: "real \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}" and w :: "'a \<Rightarrow> real" |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1535 |
and f :: "'a \<Rightarrow> 'b" and M |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1536 |
assumes "f \<in> borel_measurable M" "\<And>t. s t \<in> borel_measurable M" "integrable M w" |
61973 | 1537 |
assumes lim: "AE x in M. ((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top" |
61897
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1538 |
assumes bound: "\<forall>\<^sub>F i in at_top. AE x in M. norm (s i x) \<le> w x" |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1539 |
begin |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1540 |
|
61973 | 1541 |
lemma integral_dominated_convergence_at_top: "((\<lambda>t. integral\<^sup>L M (s t)) \<longlongrightarrow> integral\<^sup>L M f) at_top" |
61897
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1542 |
proof (rule tendsto_at_topI_sequentially) |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1543 |
fix X :: "nat \<Rightarrow> real" assume X: "filterlim X at_top sequentially" |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1544 |
from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1545 |
obtain N where w: "\<And>n. N \<le> n \<Longrightarrow> AE x in M. norm (s (X n) x) \<le> w x" |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1546 |
by (auto simp: eventually_sequentially) |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1547 |
|
61969 | 1548 |
show "(\<lambda>n. integral\<^sup>L M (s (X n))) \<longlonglongrightarrow> integral\<^sup>L M f" |
61897
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1549 |
proof (rule LIMSEQ_offset, rule integral_dominated_convergence) |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1550 |
show "AE x in M. norm (s (X (n + N)) x) \<le> w x" for n |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1551 |
by (rule w) auto |
61969 | 1552 |
show "AE x in M. (\<lambda>n. s (X (n + N)) x) \<longlonglongrightarrow> f x" |
61897
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1553 |
using lim |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1554 |
proof eventually_elim |
61973 | 1555 |
fix x assume "((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top" |
61969 | 1556 |
then show "(\<lambda>n. s (X (n + N)) x) \<longlonglongrightarrow> f x" |
61897
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1557 |
by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1558 |
qed |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1559 |
qed fact+ |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1560 |
qed |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1561 |
|
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1562 |
lemma integrable_dominated_convergence_at_top: "integrable M f" |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1563 |
proof - |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1564 |
from bound obtain N where w: "\<And>n. N \<le> n \<Longrightarrow> AE x in M. norm (s n x) \<le> w x" |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1565 |
by (auto simp: eventually_at_top_linorder) |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1566 |
show ?thesis |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1567 |
proof (rule integrable_dominated_convergence) |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1568 |
show "AE x in M. norm (s (N + i) x) \<le> w x" for i :: nat |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1569 |
by (intro w) auto |
61969 | 1570 |
show "AE x in M. (\<lambda>i. s (N + real i) x) \<longlonglongrightarrow> f x" |
61897
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1571 |
using lim |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1572 |
proof eventually_elim |
61973 | 1573 |
fix x assume "((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top" |
61969 | 1574 |
then show "(\<lambda>n. s (N + n) x) \<longlonglongrightarrow> f x" |
61897
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1575 |
by (rule filterlim_compose) |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1576 |
(auto intro!: filterlim_tendsto_add_at_top filterlim_real_sequentially) |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1577 |
qed |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1578 |
qed fact+ |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1579 |
qed |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1580 |
|
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1581 |
end |
bc0fc5499085
Bochner integral: prove dominated convergence at_top
hoelzl
parents:
61880
diff
changeset
|
1582 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1583 |
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
|
1584 |
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
|
1585 |
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
|
1586 |
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
|
1587 |
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
|
1588 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1589 |
lemma integrableI_nn_integral_finite: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1590 |
assumes [measurable]: "f \<in> borel_measurable M" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1591 |
and nonneg: "AE x in M. 0 \<le> f x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1592 |
and finite: "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1593 |
shows "integrable M f" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1594 |
proof (rule integrableI_bounded) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1595 |
have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1596 |
using nonneg by (intro nn_integral_cong_AE) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1597 |
with finite show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1598 |
by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1599 |
qed simp |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1600 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1601 |
lemma integral_nonneg_AE: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1602 |
fixes f :: "'a \<Rightarrow> real" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1603 |
assumes nonneg: "AE x in M. 0 \<le> f x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1604 |
shows "0 \<le> integral\<^sup>L M f" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1605 |
proof cases |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1606 |
assume f: "integrable M f" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1607 |
then have [measurable]: "f \<in> M \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1608 |
by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1609 |
have "(\<lambda>x. max 0 (f x)) \<in> M \<rightarrow>\<^sub>M borel" "\<And>x. 0 \<le> max 0 (f x)" "integrable M (\<lambda>x. max 0 (f x))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1610 |
using f by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1611 |
from this have "0 \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1612 |
proof (induction rule: borel_measurable_induct_real) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1613 |
case (add f g) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1614 |
then have "integrable M f" "integrable M g" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1615 |
by (auto intro!: integrable_bound[OF add.prems]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1616 |
with add show ?case |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1617 |
by (simp add: nn_integral_add) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1618 |
next |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1619 |
case (seq U) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1620 |
show ?case |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1621 |
proof (rule LIMSEQ_le_const) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1622 |
have U_le: "x \<in> space M \<Longrightarrow> U i x \<le> max 0 (f x)" for x i |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1623 |
using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1624 |
with seq nonneg show "(\<lambda>i. integral\<^sup>L M (U i)) \<longlonglongrightarrow> LINT x|M. max 0 (f x)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1625 |
by (intro integral_dominated_convergence) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1626 |
have "integrable M (U i)" for i |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1627 |
using seq.prems by (rule integrable_bound) (insert U_le seq, auto) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1628 |
with seq show "\<exists>N. \<forall>n\<ge>N. 0 \<le> integral\<^sup>L M (U n)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1629 |
by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1630 |
qed |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63627
diff
changeset
|
1631 |
qed (auto simp: integrable_mult_left_iff) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1632 |
also have "\<dots> = integral\<^sup>L M f" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1633 |
using nonneg by (auto intro!: integral_cong_AE) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1634 |
finally show ?thesis . |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1635 |
qed (simp add: not_integrable_integral_eq) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1636 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1637 |
lemma integral_nonneg[simp]: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1638 |
fixes f :: "'a \<Rightarrow> real" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1639 |
shows "(\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> 0 \<le> integral\<^sup>L M f" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1640 |
by (intro integral_nonneg_AE) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1641 |
|
56996 | 1642 |
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
|
1643 |
assumes f: "integrable M f" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1644 |
assumes nonneg: "AE x in M. 0 \<le> f x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1645 |
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
|
1646 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1647 |
{ 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
|
1648 |
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
|
1649 |
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
|
1650 |
case (set A) then show ?case |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1651 |
by (simp add: integrable_indicator_iff ennreal_indicator emeasure_eq_ennreal_measure) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1652 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1653 |
case (mult f c) then show ?case |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1654 |
by (auto simp add: integrable_mult_left_iff nn_integral_cmult ennreal_mult integral_nonneg_AE) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1655 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1656 |
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
|
1657 |
then have "integrable M f" "integrable M g" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1658 |
by (auto intro!: integrable_bound[OF add.prems]) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1659 |
with add show ?case |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1660 |
by (simp add: nn_integral_add integral_nonneg_AE) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1661 |
next |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1662 |
case (seq U) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1663 |
show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1664 |
proof (rule LIMSEQ_unique) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1665 |
have U_le_f: "x \<in> space M \<Longrightarrow> U i x \<le> f x" for x i |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1666 |
using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1667 |
have int_U: "\<And>i. integrable M (U i)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1668 |
using seq f U_le_f by (intro integrable_bound[OF f(3)]) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1669 |
from U_le_f seq have "(\<lambda>i. integral\<^sup>L M (U i)) \<longlonglongrightarrow> integral\<^sup>L M f" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1670 |
by (intro integral_dominated_convergence) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1671 |
then show "(\<lambda>i. ennreal (integral\<^sup>L M (U i))) \<longlonglongrightarrow> ennreal (integral\<^sup>L M f)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1672 |
using seq f int_U by (simp add: f integral_nonneg_AE) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1673 |
have "(\<lambda>i. \<integral>\<^sup>+ x. U i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+ x. f x \<partial>M" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1674 |
using seq U_le_f f |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1675 |
by (intro nn_integral_dominated_convergence[where w=f]) (auto simp: integrable_iff_bounded) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1676 |
then show "(\<lambda>i. \<integral>x. U i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+x. f x \<partial>M" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1677 |
using seq int_U by simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1678 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1679 |
qed } |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1680 |
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
|
1681 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1682 |
also have "\<dots> = integral\<^sup>L M f" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1683 |
using assms by (auto intro!: integral_cong_AE simp: integral_nonneg_AE) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1684 |
also have "(\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)" |
56996 | 1685 |
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
|
1686 |
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
|
1687 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1688 |
|
64008
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
1689 |
lemma nn_integral_eq_integrable: |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
1690 |
assumes f: "f \<in> M \<rightarrow>\<^sub>M borel" "AE x in M. 0 \<le> f x" and "0 \<le> x" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
1691 |
shows "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x \<longleftrightarrow> (integrable M f \<and> integral\<^sup>L M f = x)" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
1692 |
proof (safe intro!: nn_integral_eq_integral assms) |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
1693 |
assume *: "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
1694 |
with integrableI_nn_integral_finite[OF f this] nn_integral_eq_integral[of M f, OF _ f(2)] |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
1695 |
show "integrable M f" "integral\<^sup>L M f = x" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
1696 |
by (simp_all add: * assms integral_nonneg_AE) |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
1697 |
qed |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
1698 |
|
57036 | 1699 |
lemma |
1700 |
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> 'a :: {banach, second_countable_topology}" |
|
1701 |
assumes integrable[measurable]: "\<And>i. integrable M (f i)" |
|
1702 |
and summable: "AE x in M. summable (\<lambda>i. norm (f i x))" |
|
1703 |
and sums: "summable (\<lambda>i. (\<integral>x. norm (f i x) \<partial>M))" |
|
1704 |
shows integrable_suminf: "integrable M (\<lambda>x. (\<Sum>i. f i x))" (is "integrable M ?S") |
|
1705 |
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") |
|
1706 |
and integral_suminf: "(\<integral>x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>L M (f i))" |
|
1707 |
and summable_integral: "summable (\<lambda>i. integral\<^sup>L M (f i))" |
|
1708 |
proof - |
|
1709 |
have 1: "integrable M (\<lambda>x. \<Sum>i. norm (f i x))" |
|
1710 |
proof (rule integrableI_bounded) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1711 |
have "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) = (\<integral>\<^sup>+ x. (\<Sum>i. ennreal (norm (f i x))) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1712 |
apply (intro nn_integral_cong_AE) |
57036 | 1713 |
using summable |
1714 |
apply eventually_elim |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1715 |
apply (simp add: suminf_nonneg ennreal_suminf_neq_top) |
57036 | 1716 |
done |
1717 |
also have "\<dots> = (\<Sum>i. \<integral>\<^sup>+ x. norm (f i x) \<partial>M)" |
|
1718 |
by (intro nn_integral_suminf) auto |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1719 |
also have "\<dots> = (\<Sum>i. ennreal (\<integral>x. norm (f i x) \<partial>M))" |
57036 | 1720 |
by (intro arg_cong[where f=suminf] ext nn_integral_eq_integral integrable_norm integrable) auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1721 |
finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) < \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1722 |
by (simp add: sums ennreal_suminf_neq_top less_top[symmetric] integral_nonneg_AE) |
57036 | 1723 |
qed simp |
1724 |
||
61969 | 1725 |
have 2: "AE x in M. (\<lambda>n. \<Sum>i<n. f i x) \<longlonglongrightarrow> (\<Sum>i. f i x)" |
57036 | 1726 |
using summable by eventually_elim (auto intro: summable_LIMSEQ summable_norm_cancel) |
1727 |
||
1728 |
have 3: "\<And>j. AE x in M. norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))" |
|
1729 |
using summable |
|
1730 |
proof eventually_elim |
|
1731 |
fix j x assume [simp]: "summable (\<lambda>i. norm (f i x))" |
|
64267 | 1732 |
have "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i<j. norm (f i x))" by (rule norm_sum) |
57036 | 1733 |
also have "\<dots> \<le> (\<Sum>i. norm (f i x))" |
64267 | 1734 |
using sum_le_suminf[of "\<lambda>i. norm (f i x)"] unfolding sums_iff by auto |
57036 | 1735 |
finally show "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))" by simp |
1736 |
qed |
|
1737 |
||
57137 | 1738 |
note ibl = integrable_dominated_convergence[OF _ _ 1 2 3] |
1739 |
note int = integral_dominated_convergence[OF _ _ 1 2 3] |
|
57036 | 1740 |
|
1741 |
show "integrable M ?S" |
|
57137 | 1742 |
by (rule ibl) measurable |
57036 | 1743 |
|
1744 |
show "?f sums ?x" unfolding sums_def |
|
57137 | 1745 |
using int by (simp add: integrable) |
57036 | 1746 |
then show "?x = suminf ?f" "summable ?f" |
1747 |
unfolding sums_iff by auto |
|
1748 |
qed |
|
1749 |
||
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1750 |
lemma integral_norm_bound [simp]: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1751 |
fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}" |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1752 |
shows "norm (integral\<^sup>L M f) \<le> (\<integral>x. norm (f x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1753 |
proof (cases "integrable M f") |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1754 |
case True then show ?thesis |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1755 |
using nn_integral_eq_integral[of M "\<lambda>x. norm (f x)"] integral_norm_bound_ennreal[of M f] |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1756 |
by (simp add: integral_nonneg_AE) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1757 |
next |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1758 |
case False |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1759 |
then have "norm (integral\<^sup>L M f) = 0" by (simp add: not_integrable_integral_eq) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1760 |
moreover have "(\<integral>x. norm (f x) \<partial>M) \<ge> 0" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1761 |
ultimately show ?thesis by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1762 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1763 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1764 |
lemma integral_abs_bound [simp]: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1765 |
fixes f :: "'a \<Rightarrow> real" shows "abs (\<integral>x. f x \<partial>M) \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1766 |
using integral_norm_bound[of M f] by auto |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1767 |
|
56996 | 1768 |
lemma integral_eq_nn_integral: |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1769 |
assumes [measurable]: "f \<in> borel_measurable M" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1770 |
assumes nonneg: "AE x in M. 0 \<le> f x" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1771 |
shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)" |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1772 |
proof cases |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1773 |
assume *: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1774 |
also have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1775 |
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
|
1776 |
finally have "\<not> integrable M f" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1777 |
by (auto simp: integrable_iff_bounded) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1778 |
then show ?thesis |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1779 |
by (simp add: * not_integrable_integral_eq) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1780 |
next |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1781 |
assume "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1782 |
then have "integrable M f" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1783 |
by (cases "\<integral>\<^sup>+ x. ennreal (f x) \<partial>M" rule: ennreal_cases) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1784 |
(auto intro!: integrableI_nn_integral_finite assms) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1785 |
from nn_integral_eq_integral[OF this] nonneg show ?thesis |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1786 |
by (simp add: integral_nonneg_AE) |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1787 |
qed |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1788 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1789 |
lemma enn2real_nn_integral_eq_integral: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1790 |
assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \<le> g x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1791 |
and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < top" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1792 |
and [measurable]: "g \<in> M \<rightarrow>\<^sub>M borel" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1793 |
shows "enn2real (\<integral>\<^sup>+x. f x \<partial>M) = (\<integral>x. g x \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1794 |
proof - |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1795 |
have "ennreal (enn2real (\<integral>\<^sup>+x. f x \<partial>M)) = (\<integral>\<^sup>+x. f x \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1796 |
using fin by (intro ennreal_enn2real) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1797 |
also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1798 |
using eq by (rule nn_integral_cong_AE) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1799 |
also have "\<dots> = (\<integral>x. g x \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1800 |
proof (rule nn_integral_eq_integral) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1801 |
show "integrable M g" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1802 |
proof (rule integrableI_bounded) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1803 |
have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1804 |
using eq nn by (auto intro!: nn_integral_cong_AE elim!: eventually_elim2) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1805 |
also note fin |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1806 |
finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1807 |
by simp |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1808 |
qed simp |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1809 |
qed fact |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1810 |
finally show ?thesis |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1811 |
using nn by (simp add: integral_nonneg_AE) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1812 |
qed |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1813 |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1814 |
lemma has_bochner_integral_nn_integral: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1815 |
assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1816 |
assumes "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x" |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1817 |
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
|
1818 |
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
|
1819 |
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
|
1820 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1821 |
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
|
1822 |
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
|
1823 |
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
|
1824 |
by (intro integrableI_sequence[where s="\<lambda>_. f"] borel_measurable_simple_function) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1825 |
(auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1826 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1827 |
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
|
1828 |
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
|
1829 |
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
|
1830 |
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
|
1831 |
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
|
1832 |
assumes lim: "\<And>f s. (\<And>i. integrable M (s i)) \<Longrightarrow> (\<And>i. P (s i)) \<Longrightarrow> |
61969 | 1833 |
(\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x) \<Longrightarrow> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1834 |
(\<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
|
1835 |
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
|
1836 |
proof - |
61808 | 1837 |
from \<open>integrable M f\<close> have f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1838 |
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
|
1839 |
from borel_measurable_implies_sequence_metric[OF f(1)] |
61969 | 1840 |
obtain s where s: "\<And>i. simple_function M (s i)" "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1841 |
"\<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
|
1842 |
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
|
1843 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1844 |
{ fix f A |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1845 |
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
|
1846 |
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
|
1847 |
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
|
1848 |
(\<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
|
1849 |
by (induct A rule: infinite_finite_induct) (auto intro!: add) } |
64267 | 1850 |
note sum = this |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1851 |
|
63040 | 1852 |
define s' where [abs_def]: "s' i z = indicator (space M) z *\<^sub>R s i z" for i z |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1853 |
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
|
1854 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1855 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1856 |
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
|
1857 |
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
|
1858 |
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
|
1859 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1860 |
{ fix i |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1861 |
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
|
1862 |
(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
|
1863 |
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
|
1864 |
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
|
1865 |
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
|
1866 |
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
|
1867 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1868 |
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
|
1869 |
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
|
1870 |
fix i |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1871 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1872 |
have "(\<integral>\<^sup>+x. norm (s' i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)" |
56996 | 1873 |
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
|
1874 |
also have "\<dots> < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1875 |
using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1876 |
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
|
1877 |
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
|
1878 |
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
|
1879 |
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
|
1880 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1881 |
{ 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
|
1882 |
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
|
1883 |
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
|
1884 |
also have "\<dots> < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1885 |
using sbi by (auto elim: simple_bochner_integrable.cases simp: less_top) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1886 |
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
|
1887 |
then show "P (s' i)" |
64267 | 1888 |
by (subst s'_eq) (auto intro!: sum base simp: less_top) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1889 |
|
61969 | 1890 |
fix x assume "x \<in> space M" with s show "(\<lambda>i. s' i x) \<longlonglongrightarrow> f x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1891 |
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
|
1892 |
show "norm (s' i x) \<le> 2 * norm (f x)" |
61808 | 1893 |
using \<open>x \<in> space M\<close> s by (simp add: s'_eq_s) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1894 |
qed fact |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1895 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1896 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1897 |
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
|
1898 |
"(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
|
1899 |
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
|
1900 |
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
|
1901 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1902 |
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
|
1903 |
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
|
1904 |
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
|
1905 |
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
|
1906 |
proof |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1907 |
assume "integral\<^sup>L M f = 0" |
56996 | 1908 |
then have "integral\<^sup>N M f = 0" |
1909 |
using nn_integral_eq_integral[OF f nonneg] by simp |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1910 |
then have "AE x in M. ennreal (f x) \<le> 0" |
56996 | 1911 |
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
|
1912 |
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
|
1913 |
by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1914 |
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
|
1915 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1916 |
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
|
1917 |
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
|
1918 |
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
|
1919 |
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
|
1920 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1921 |
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
|
1922 |
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
|
1923 |
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
|
1924 |
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
|
1925 |
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
|
1926 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1927 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1928 |
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
|
1929 |
fixes f :: "'a \<Rightarrow> real" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1930 |
shows "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x \<le> g x) \<Longrightarrow> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
1931 |
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
|
1932 |
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
|
1933 |
|
64911 | 1934 |
text \<open>The next two statements are useful to bound Lebesgue integrals, as they avoid one |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1935 |
integrability assumption. The price to pay is that the upper function has to be nonnegative, |
64911 | 1936 |
but this is often true and easy to check in computations.\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1937 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1938 |
lemma integral_mono_AE': |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1939 |
fixes f::"_ \<Rightarrow> real" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1940 |
assumes "integrable M f" "AE x in M. g x \<le> f x" "AE x in M. 0 \<le> f x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1941 |
shows "(\<integral>x. g x \<partial>M) \<le> (\<integral>x. f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1942 |
proof (cases "integrable M g") |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1943 |
case True |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1944 |
show ?thesis by (rule integral_mono_AE, auto simp add: assms True) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1945 |
next |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1946 |
case False |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1947 |
then have "(\<integral>x. g x \<partial>M) = 0" by (simp add: not_integrable_integral_eq) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1948 |
also have "... \<le> (\<integral>x. f x \<partial>M)" by (simp add: integral_nonneg_AE[OF assms(3)]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1949 |
finally show ?thesis by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1950 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1951 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1952 |
lemma integral_mono': |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1953 |
fixes f::"_ \<Rightarrow> real" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1954 |
assumes "integrable M f" "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1955 |
shows "(\<integral>x. g x \<partial>M) \<le> (\<integral>x. f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1956 |
by (rule integral_mono_AE', insert assms, auto) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1957 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1958 |
lemma (in finite_measure) integrable_measure: |
57025 | 1959 |
assumes I: "disjoint_family_on X I" "countable I" |
1960 |
shows "integrable (count_space I) (\<lambda>i. measure M (X i))" |
|
1961 |
proof - |
|
1962 |
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)" |
|
1963 |
by (auto intro!: nn_integral_cong measure_notin_sets) |
|
1964 |
also have "\<dots> = measure M (\<Union>i\<in>I. if X i \<in> sets M then X i else {})" |
|
1965 |
using I unfolding emeasure_eq_measure[symmetric] |
|
1966 |
by (subst emeasure_UN_countable) (auto simp: disjoint_family_on_def) |
|
1967 |
finally show ?thesis |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1968 |
by (auto intro!: integrableI_bounded) |
57025 | 1969 |
qed |
1970 |
||
1971 |
lemma integrableI_real_bounded: |
|
1972 |
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>" |
|
1973 |
shows "integrable M f" |
|
1974 |
proof (rule integrableI_bounded) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1975 |
have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = \<integral>\<^sup>+ x. ennreal (f x) \<partial>M" |
57025 | 1976 |
using ae by (auto intro: nn_integral_cong_AE) |
1977 |
also note fin |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1978 |
finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" . |
57025 | 1979 |
qed fact |
1980 |
||
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1981 |
lemma nn_integral_nonneg_infinite: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1982 |
fixes f::"'a \<Rightarrow> real" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1983 |
assumes "f \<in> borel_measurable M" "\<not> integrable M f" "AE x in M. f x \<ge> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1984 |
shows "(\<integral>\<^sup>+x. f x \<partial>M) = \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1985 |
using assms integrableI_real_bounded less_top by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
1986 |
|
57025 | 1987 |
lemma integral_real_bounded: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1988 |
assumes "0 \<le> r" "integral\<^sup>N M f \<le> ennreal r" |
57025 | 1989 |
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
|
1990 |
proof cases |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1991 |
assume [simp]: "integrable M f" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1992 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1993 |
have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = integral\<^sup>N M (\<lambda>x. max 0 (f x))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1994 |
by (intro nn_integral_eq_integral[symmetric]) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1995 |
also have "\<dots> = integral\<^sup>N M f" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1996 |
by (intro nn_integral_cong) (simp add: max_def ennreal_neg) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1997 |
also have "\<dots> \<le> r" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1998 |
by fact |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
1999 |
finally have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) \<le> r" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2000 |
using \<open>0 \<le> r\<close> by simp |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2001 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2002 |
moreover have "integral\<^sup>L M f \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2003 |
by (rule integral_mono_AE) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2004 |
ultimately show ?thesis |
57025 | 2005 |
by simp |
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
2006 |
next |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2007 |
assume "\<not> integrable M f" then show ?thesis |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2008 |
using \<open>0 \<le> r\<close> by (simp add: not_integrable_integral_eq) |
57025 | 2009 |
qed |
2010 |
||
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2011 |
lemma integrable_Min: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2012 |
fixes f:: "_ \<Rightarrow> _ \<Rightarrow> real" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2013 |
assumes "finite I" "I \<noteq> {}" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2014 |
"\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2015 |
shows "integrable M (\<lambda>x. Min {f i x|i. i \<in> I})" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2016 |
using assms apply (induct rule: finite_ne_induct) apply simp+ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2017 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2018 |
fix j F assume H: "finite F" "F \<noteq> {}" "integrable M (\<lambda>x. Min {f i x |i. i \<in> F})" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2019 |
"(\<And>i. i = j \<or> i \<in> F \<Longrightarrow> integrable M (f i))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2020 |
have "{f i x |i. i = j \<or> i \<in> F} = insert (f j x) {f i x |i. i \<in> F}" for x by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2021 |
then have "Min {f i x |i. i = j \<or> i \<in> F} = min (f j x) (Min {f i x |i. i \<in> F})" for x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2022 |
using H(1) H(2) Min_insert by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2023 |
moreover have "integrable M (\<lambda>x. min (f j x) (Min {f i x |i. i \<in> F}))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2024 |
by (rule integrable_min, auto simp add: H) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2025 |
ultimately show "integrable M (\<lambda>x. Min {f i x |i. i = j \<or> i \<in> F})" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2026 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2027 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2028 |
lemma integrable_Max: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2029 |
fixes f:: "_ \<Rightarrow> _ \<Rightarrow> real" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2030 |
assumes "finite I" "I \<noteq> {}" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2031 |
"\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2032 |
shows "integrable M (\<lambda>x. Max {f i x|i. i \<in> I})" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2033 |
using assms apply (induct rule: finite_ne_induct) apply simp+ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2034 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2035 |
fix j F assume H: "finite F" "F \<noteq> {}" "integrable M (\<lambda>x. Max {f i x |i. i \<in> F})" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2036 |
"(\<And>i. i = j \<or> i \<in> F \<Longrightarrow> integrable M (f i))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2037 |
have "{f i x |i. i = j \<or> i \<in> F} = insert (f j x) {f i x |i. i \<in> F}" for x by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2038 |
then have "Max {f i x |i. i = j \<or> i \<in> F} = max (f j x) (Max {f i x |i. i \<in> F})" for x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2039 |
using H(1) H(2) Max_insert by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2040 |
moreover have "integrable M (\<lambda>x. max (f j x) (Max {f i x |i. i \<in> F}))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2041 |
by (rule integrable_max, auto simp add: H) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2042 |
ultimately show "integrable M (\<lambda>x. Max {f i x |i. i = j \<or> i \<in> F})" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2043 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2044 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2045 |
lemma integral_Markov_inequality: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2046 |
assumes [measurable]: "integrable M u" and "AE x in M. 0 \<le> u x" "0 < (c::real)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2047 |
shows "(emeasure M) {x\<in>space M. u x \<ge> c} \<le> (1/c) * (\<integral>x. u x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2048 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2049 |
have "(\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M) \<le> (\<integral>\<^sup>+ x. u x \<partial>M)" |
64911 | 2050 |
by (rule nn_integral_mono_AE, auto simp add: \<open>c>0\<close> less_eq_real_def) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2051 |
also have "... = (\<integral>x. u x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2052 |
by (rule nn_integral_eq_integral, auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2053 |
finally have *: "(\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M) \<le> (\<integral>x. u x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2054 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2055 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2056 |
have "{x \<in> space M. u x \<ge> c} = {x \<in> space M. ennreal(1/c) * u x \<ge> 1} \<inter> (space M)" |
64911 | 2057 |
using \<open>c>0\<close> by (auto simp: ennreal_mult'[symmetric]) |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2058 |
then have "emeasure M {x \<in> space M. u x \<ge> c} = emeasure M ({x \<in> space M. ennreal(1/c) * u x \<ge> 1} \<inter> (space M))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2059 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2060 |
also have "... \<le> ennreal(1/c) * (\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2061 |
by (rule nn_integral_Markov_inequality) (auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2062 |
also have "... \<le> ennreal(1/c) * (\<integral>x. u x \<partial>M)" |
64911 | 2063 |
apply (rule mult_left_mono) using * \<open>c>0\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2064 |
finally show ?thesis |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2065 |
using \<open>0<c\<close> by (simp add: ennreal_mult'[symmetric]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2066 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2067 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2068 |
lemma integral_ineq_eq_0_then_AE: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2069 |
fixes f::"_ \<Rightarrow> real" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2070 |
assumes "AE x in M. f x \<le> g x" "integrable M f" "integrable M g" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2071 |
"(\<integral>x. f x \<partial>M) = (\<integral>x. g x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2072 |
shows "AE x in M. f x = g x" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2073 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2074 |
define h where "h = (\<lambda>x. g x - f x)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2075 |
have "AE x in M. h x = 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2076 |
apply (subst integral_nonneg_eq_0_iff_AE[symmetric]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2077 |
unfolding h_def using assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2078 |
then show ?thesis unfolding h_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2079 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2080 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2081 |
lemma not_AE_zero_int_E: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2082 |
fixes f::"'a \<Rightarrow> real" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2083 |
assumes "AE x in M. f x \<ge> 0" "(\<integral>x. f x \<partial>M) > 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2084 |
and [measurable]: "f \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2085 |
shows "\<exists>A e. A \<in> sets M \<and> e>0 \<and> emeasure M A > 0 \<and> (\<forall>x \<in> A. f x \<ge> e)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2086 |
proof (rule not_AE_zero_E, auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2087 |
assume *: "AE x in M. f x = 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2088 |
have "(\<integral>x. f x \<partial>M) = (\<integral>x. 0 \<partial>M)" by (rule integral_cong_AE, auto simp add: *) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2089 |
then have "(\<integral>x. f x \<partial>M) = 0" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2090 |
then show False using assms(2) by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2091 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2092 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2093 |
lemma tendsto_L1_int: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2094 |
fixes u :: "_ \<Rightarrow> _ \<Rightarrow> 'b::{banach, second_countable_topology}" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2095 |
assumes [measurable]: "\<And>n. integrable M (u n)" "integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2096 |
and "((\<lambda>n. (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)) \<longlongrightarrow> 0) F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2097 |
shows "((\<lambda>n. (\<integral>x. u n x \<partial>M)) \<longlongrightarrow> (\<integral>x. f x \<partial>M)) F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2098 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2099 |
have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> (0::ennreal)) F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2100 |
proof (rule tendsto_sandwich[of "\<lambda>_. 0", where ?h = "\<lambda>n. (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)"], auto simp add: assms) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2101 |
{ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2102 |
fix n |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2103 |
have "(\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M) = (\<integral>x. u n x - f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2104 |
apply (rule Bochner_Integration.integral_diff[symmetric]) using assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2105 |
then have "norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) = norm (\<integral>x. u n x - f x \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2106 |
by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2107 |
also have "... \<le> (\<integral>x. norm(u n x - f x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2108 |
by (rule integral_norm_bound) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2109 |
finally have "ennreal(norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<le> (\<integral>x. norm(u n x - f x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2110 |
by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2111 |
also have "... = (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2112 |
apply (rule nn_integral_eq_integral[symmetric]) using assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2113 |
finally have "norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2114 |
} |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2115 |
then show "eventually (\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)) F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2116 |
by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2117 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2118 |
then have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2119 |
by (simp add: ennreal_0[symmetric] del: ennreal_0) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2120 |
then have "((\<lambda>n. ((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F" using tendsto_norm_zero_iff by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2121 |
then show ?thesis using Lim_null by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2122 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2123 |
|
64911 | 2124 |
text \<open>The next lemma asserts that, if a sequence of functions converges in $L^1$, then |
2125 |
it admits a subsequence that converges almost everywhere.\<close> |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2126 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2127 |
lemma tendsto_L1_AE_subseq: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2128 |
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2129 |
assumes [measurable]: "\<And>n. integrable M (u n)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2130 |
and "(\<lambda>n. (\<integral>x. norm(u n x) \<partial>M)) \<longlonglongrightarrow> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2131 |
shows "\<exists>r. subseq r \<and> (AE x in M. (\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2132 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2133 |
{ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2134 |
fix k |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2135 |
have "eventually (\<lambda>n. (\<integral>x. norm(u n x) \<partial>M) < (1/2)^k) sequentially" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2136 |
using order_tendstoD(2)[OF assms(2)] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2137 |
with eventually_elim2[OF eventually_gt_at_top[of k] this] |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2138 |
have "\<exists>n>k. (\<integral>x. norm(u n x) \<partial>M) < (1/2)^k" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2139 |
by (metis eventually_False_sequentially) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2140 |
} |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2141 |
then have "\<exists>r. \<forall>n. True \<and> (r (Suc n) > r n \<and> (\<integral>x. norm(u (r (Suc n)) x) \<partial>M) < (1/2)^(r n))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2142 |
by (intro dependent_nat_choice, auto) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2143 |
then obtain r0 where r0: "subseq r0" "\<And>n. (\<integral>x. norm(u (r0 (Suc n)) x) \<partial>M) < (1/2)^(r0 n)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2144 |
by (auto simp: subseq_Suc_iff) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2145 |
define r where "r = (\<lambda>n. r0(n+1))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2146 |
have "subseq r" unfolding r_def using r0(1) by (simp add: subseq_Suc_iff) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2147 |
have I: "(\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M) < ennreal((1/2)^n)" for n |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2148 |
proof - |
64911 | 2149 |
have "r0 n \<ge> n" using \<open>subseq r0\<close> by (simp add: seq_suble) |
2150 |
have "(1/2::real)^(r0 n) \<le> (1/2)^n" by (rule power_decreasing[OF \<open>r0 n \<ge> n\<close>], auto) |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2151 |
then have "(\<integral>x. norm(u (r0 (Suc n)) x) \<partial>M) < (1/2)^n" using r0(2) less_le_trans by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2152 |
then have "(\<integral>x. norm(u (r n) x) \<partial>M) < (1/2)^n" unfolding r_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2153 |
moreover have "(\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M) = (\<integral>x. norm(u (r n) x) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2154 |
by (rule nn_integral_eq_integral, auto simp add: integrable_norm[OF assms(1)[of "r n"]]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2155 |
ultimately show ?thesis by (auto intro: ennreal_lessI) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2156 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2157 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2158 |
have "AE x in M. limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2159 |
proof (rule AE_upper_bound_inf_ennreal) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2160 |
fix e::real assume "e > 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2161 |
define A where "A = (\<lambda>n. {x \<in> space M. norm(u (r n) x) \<ge> e})" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2162 |
have A_meas [measurable]: "\<And>n. A n \<in> sets M" unfolding A_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2163 |
have A_bound: "emeasure M (A n) < (1/e) * ennreal((1/2)^n)" for n |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2164 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2165 |
have *: "indicator (A n) x \<le> (1/e) * ennreal(norm(u (r n) x))" for x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2166 |
apply (cases "x \<in> A n") unfolding A_def using \<open>0 < e\<close> by (auto simp: ennreal_mult[symmetric]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2167 |
have "emeasure M (A n) = (\<integral>\<^sup>+x. indicator (A n) x \<partial>M)" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2168 |
also have "... \<le> (\<integral>\<^sup>+x. (1/e) * ennreal(norm(u (r n) x)) \<partial>M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2169 |
apply (rule nn_integral_mono) using * by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2170 |
also have "... = (1/e) * (\<integral>\<^sup>+x. norm(u (r n) x) \<partial>M)" |
64911 | 2171 |
apply (rule nn_integral_cmult) using \<open>e > 0\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2172 |
also have "... < (1/e) * ennreal((1/2)^n)" |
64911 | 2173 |
using I[of n] \<open>e > 0\<close> by (intro ennreal_mult_strict_left_mono) auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2174 |
finally show ?thesis by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2175 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2176 |
have A_fin: "emeasure M (A n) < \<infinity>" for n |
64911 | 2177 |
using \<open>e > 0\<close> A_bound[of n] |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2178 |
by (auto simp add: ennreal_mult less_top[symmetric]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2179 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2180 |
have A_sum: "summable (\<lambda>n. measure M (A n))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2181 |
proof (rule summable_comparison_test'[of "\<lambda>n. (1/e) * (1/2)^n" 0]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2182 |
have "summable (\<lambda>n. (1/(2::real))^n)" by (simp add: summable_geometric) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2183 |
then show "summable (\<lambda>n. (1/e) * (1/2)^n)" using summable_mult by blast |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2184 |
fix n::nat assume "n \<ge> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2185 |
have "norm(measure M (A n)) = measure M (A n)" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2186 |
also have "... = enn2real(emeasure M (A n))" unfolding measure_def by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2187 |
also have "... < enn2real((1/e) * (1/2)^n)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2188 |
using A_bound[of n] \<open>emeasure M (A n) < \<infinity>\<close> \<open>0 < e\<close> |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2189 |
by (auto simp: emeasure_eq_ennreal_measure ennreal_mult[symmetric] ennreal_less_iff) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2190 |
also have "... = (1/e) * (1/2)^n" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2191 |
using \<open>0<e\<close> by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2192 |
finally show "norm(measure M (A n)) \<le> (1/e) * (1/2)^n" by simp |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2193 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2194 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2195 |
have "AE x in M. eventually (\<lambda>n. x \<in> space M - A n) sequentially" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2196 |
by (rule borel_cantelli_AE1[OF A_meas A_fin A_sum]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2197 |
moreover |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2198 |
{ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2199 |
fix x assume "eventually (\<lambda>n. x \<in> space M - A n) sequentially" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2200 |
moreover have "norm(u (r n) x) \<le> ennreal e" if "x \<in> space M - A n" for n |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2201 |
using that unfolding A_def by (auto intro: ennreal_leI) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2202 |
ultimately have "eventually (\<lambda>n. norm(u (r n) x) \<le> ennreal e) sequentially" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2203 |
by (simp add: eventually_mono) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2204 |
then have "limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> e" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2205 |
by (simp add: Limsup_bounded) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2206 |
} |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2207 |
ultimately show "AE x in M. limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0 + ennreal e" by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2208 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2209 |
moreover |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2210 |
{ |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2211 |
fix x assume "limsup (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2212 |
moreover then have "liminf (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2213 |
by (rule order_trans[rotated]) (auto intro: Liminf_le_Limsup) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2214 |
ultimately have "(\<lambda>n. ennreal (norm(u (r n) x))) \<longlonglongrightarrow> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2215 |
using tendsto_Limsup[of sequentially "\<lambda>n. ennreal (norm(u (r n) x))"] by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2216 |
then have "(\<lambda>n. norm(u (r n) x)) \<longlonglongrightarrow> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2217 |
by (simp add: ennreal_0[symmetric] del: ennreal_0) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2218 |
then have "(\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2219 |
by (simp add: tendsto_norm_zero_iff) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2220 |
} |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2221 |
ultimately have "AE x in M. (\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0" by auto |
64911 | 2222 |
then show ?thesis using \<open>subseq r\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2223 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2224 |
|
61808 | 2225 |
subsection \<open>Restricted measure spaces\<close> |
57137 | 2226 |
|
2227 |
lemma integrable_restrict_space: |
|
2228 |
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
|
2229 |
assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M" |
|
2230 |
shows "integrable (restrict_space M \<Omega>) f \<longleftrightarrow> integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" |
|
2231 |
unfolding integrable_iff_bounded |
|
2232 |
borel_measurable_restrict_space_iff[OF \<Omega>] |
|
2233 |
nn_integral_restrict_space[OF \<Omega>] |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2234 |
by (simp add: ac_simps ennreal_indicator ennreal_mult) |
57137 | 2235 |
|
2236 |
lemma integral_restrict_space: |
|
2237 |
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
|
2238 |
assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M" |
|
2239 |
shows "integral\<^sup>L (restrict_space M \<Omega>) f = integral\<^sup>L M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" |
|
2240 |
proof (rule integral_eq_cases) |
|
2241 |
assume "integrable (restrict_space M \<Omega>) f" |
|
2242 |
then show ?thesis |
|
2243 |
proof induct |
|
2244 |
case (base A c) then show ?case |
|
2245 |
by (simp add: indicator_inter_arith[symmetric] sets_restrict_space_iff |
|
2246 |
emeasure_restrict_space Int_absorb1 measure_restrict_space) |
|
2247 |
next |
|
2248 |
case (add g f) then show ?case |
|
2249 |
by (simp add: scaleR_add_right integrable_restrict_space) |
|
2250 |
next |
|
2251 |
case (lim f s) |
|
2252 |
show ?case |
|
2253 |
proof (rule LIMSEQ_unique) |
|
61969 | 2254 |
show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) \<longlonglongrightarrow> integral\<^sup>L (restrict_space M \<Omega>) f" |
57137 | 2255 |
using lim by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) simp_all |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2256 |
|
61969 | 2257 |
show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) \<longlonglongrightarrow> (\<integral> x. indicator \<Omega> x *\<^sub>R f x \<partial>M)" |
57137 | 2258 |
unfolding lim |
2259 |
using lim |
|
2260 |
by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (indicator \<Omega> x *\<^sub>R f x)"]) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2261 |
(auto simp add: space_restrict_space integrable_restrict_space simp del: norm_scaleR |
57137 | 2262 |
split: split_indicator) |
2263 |
qed |
|
2264 |
qed |
|
2265 |
qed (simp add: integrable_restrict_space) |
|
2266 |
||
60066 | 2267 |
lemma integral_empty: |
2268 |
assumes "space M = {}" |
|
2269 |
shows "integral\<^sup>L M f = 0" |
|
2270 |
proof - |
|
2271 |
have "(\<integral> x. f x \<partial>M) = (\<integral> x. 0 \<partial>M)" |
|
2272 |
by(rule integral_cong)(simp_all add: assms) |
|
2273 |
thus ?thesis by simp |
|
2274 |
qed |
|
2275 |
||
61808 | 2276 |
subsection \<open>Measure spaces with an associated density\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2277 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2278 |
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
|
2279 |
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
|
2280 |
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
|
2281 |
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
|
2282 |
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
|
2283 |
unfolding integrable_iff_bounded using nn |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2284 |
apply (simp add: nn_integral_density less_top[symmetric]) |
56996 | 2285 |
apply (intro arg_cong2[where f="op ="] refl nn_integral_cong_AE) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2286 |
apply (auto simp: ennreal_mult) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2287 |
done |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2288 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2289 |
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
|
2290 |
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
|
2291 |
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
|
2292 |
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
|
2293 |
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
|
2294 |
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
|
2295 |
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
|
2296 |
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
|
2297 |
proof induct |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2298 |
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
|
2299 |
then have [measurable]: "A \<in> sets M" by auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2300 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2301 |
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
|
2302 |
using g base integrable_density[of "indicator A :: 'a \<Rightarrow> real" M g] by simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2303 |
then have "integral\<^sup>L M (\<lambda>x. g x * indicator A x) = (\<integral>\<^sup>+ x. ennreal (g x * indicator A x) \<partial>M)" |
56996 | 2304 |
using g by (subst nn_integral_eq_integral) auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2305 |
also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (g x) * indicator A x \<partial>M)" |
56996 | 2306 |
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
|
2307 |
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
|
2308 |
by (rule emeasure_density[symmetric]) auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2309 |
also have "\<dots> = ennreal (measure (density M g) A)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2310 |
using base by (auto intro: emeasure_eq_ennreal_measure) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2311 |
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
|
2312 |
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
|
2313 |
finally show ?case |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2314 |
using base g |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2315 |
apply (simp add: int integral_nonneg_AE) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2316 |
apply (subst (asm) ennreal_inj) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2317 |
apply (auto intro!: integral_nonneg_AE) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2318 |
done |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2319 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2320 |
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
|
2321 |
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
|
2322 |
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
|
2323 |
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
|
2324 |
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
|
2325 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2326 |
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
|
2327 |
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
|
2328 |
using lim(1,5)[THEN borel_measurable_integrable] by auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2329 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2330 |
show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2331 |
proof (rule LIMSEQ_unique) |
61969 | 2332 |
show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. g x *\<^sub>R f x)" |
57137 | 2333 |
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
|
2334 |
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
|
2335 |
by (intro integrable_mult_right integrable_norm integrable_density[THEN iffD1] lim g) auto |
61969 | 2336 |
show "AE x in M. (\<lambda>i. g x *\<^sub>R s i x) \<longlonglongrightarrow> g x *\<^sub>R f x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2337 |
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
|
2338 |
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
|
2339 |
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
|
2340 |
qed auto |
61969 | 2341 |
show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) \<longlonglongrightarrow> integral\<^sup>L (density M g) f" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2342 |
unfolding lim(2)[symmetric] |
57137 | 2343 |
by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) |
57036 | 2344 |
(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
|
2345 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2346 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2347 |
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
|
2348 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2349 |
lemma |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2350 |
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
|
2351 |
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
|
2352 |
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
|
2353 |
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
|
2354 |
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
|
2355 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2356 |
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
|
2357 |
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
|
2358 |
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
|
2359 |
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
|
2360 |
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
|
2361 |
|
61808 | 2362 |
subsection \<open>Distributions\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2363 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2364 |
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
|
2365 |
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
|
2366 |
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
|
2367 |
shows "integrable (distr M N g) f \<longleftrightarrow> integrable M (\<lambda>x. f (g x))" |
56996 | 2368 |
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
|
2369 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2370 |
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
|
2371 |
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
|
2372 |
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
|
2373 |
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
|
2374 |
(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
|
2375 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2376 |
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
|
2377 |
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
|
2378 |
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
|
2379 |
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
|
2380 |
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
|
2381 |
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
|
2382 |
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
|
2383 |
proof induct |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2384 |
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
|
2385 |
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
|
2386 |
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
|
2387 |
by (intro integrable_indicator) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2388 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2389 |
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
|
2390 |
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
|
2391 |
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
|
2392 |
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
|
2393 |
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
|
2394 |
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
|
2395 |
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
|
2396 |
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
|
2397 |
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
|
2398 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2399 |
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
|
2400 |
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
|
2401 |
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
|
2402 |
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
|
2403 |
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
|
2404 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2405 |
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
|
2406 |
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
|
2407 |
using lim(1,5)[THEN borel_measurable_integrable] by auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2408 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2409 |
show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2410 |
proof (rule LIMSEQ_unique) |
61969 | 2411 |
show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. f (g x))" |
57137 | 2412 |
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
|
2413 |
show "integrable M (\<lambda>x. 2 * norm (f (g x)))" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2414 |
using lim by (auto simp: integrable_distr_eq) |
61969 | 2415 |
show "AE x in M. (\<lambda>i. s i (g x)) \<longlonglongrightarrow> f (g x)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2416 |
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
|
2417 |
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
|
2418 |
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
|
2419 |
qed auto |
61969 | 2420 |
show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) \<longlonglongrightarrow> integral\<^sup>L (distr M N g) f" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2421 |
unfolding lim(2)[symmetric] |
57137 | 2422 |
by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) |
57036 | 2423 |
(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
|
2424 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2425 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2426 |
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
|
2427 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2428 |
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
|
2429 |
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
|
2430 |
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
|
2431 |
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
|
2432 |
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
|
2433 |
|
61808 | 2434 |
subsection \<open>Lebesgue integration on @{const count_space}\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2435 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2436 |
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
|
2437 |
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
|
2438 |
shows "finite X \<Longrightarrow> integrable (count_space X) f" |
56996 | 2439 |
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
|
2440 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2441 |
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
|
2442 |
"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
|
2443 |
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
|
2444 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2445 |
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
|
2446 |
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
|
2447 |
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
|
2448 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2449 |
have eq: "\<And>x. x \<in> A \<Longrightarrow> (\<Sum>a | x = a \<and> a \<in> A \<and> f a \<noteq> 0. f a) = (\<Sum>x\<in>{x}. f x)" |
64267 | 2450 |
by (intro sum.mono_neutral_cong_left) auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2451 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2452 |
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
|
2453 |
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
|
2454 |
also have "\<dots> = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. measure (count_space A) {a} *\<^sub>R f a)" |
64267 | 2455 |
by (subst integral_sum) (auto intro!: sum.cong) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2456 |
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
|
2457 |
by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2458 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2459 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2460 |
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
|
2461 |
by (subst lebesgue_integral_count_space_finite_support) |
64267 | 2462 |
(auto intro!: sum.mono_neutral_cong_left) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2463 |
|
59425 | 2464 |
lemma integrable_count_space_nat_iff: |
2465 |
fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}" |
|
2466 |
shows "integrable (count_space UNIV) f \<longleftrightarrow> summable (\<lambda>x. norm (f x))" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2467 |
by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2468 |
intro: summable_suminf_not_top) |
59425 | 2469 |
|
2470 |
lemma sums_integral_count_space_nat: |
|
2471 |
fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}" |
|
2472 |
assumes *: "integrable (count_space UNIV) f" |
|
2473 |
shows "f sums (integral\<^sup>L (count_space UNIV) f)" |
|
2474 |
proof - |
|
2475 |
let ?f = "\<lambda>n i. indicator {n} i *\<^sub>R f i" |
|
2476 |
have f': "\<And>n i. ?f n i = indicator {n} i *\<^sub>R f n" |
|
2477 |
by (auto simp: fun_eq_iff split: split_indicator) |
|
2478 |
||
2479 |
have "(\<lambda>i. \<integral>n. ?f i n \<partial>count_space UNIV) sums \<integral> n. (\<Sum>i. ?f i n) \<partial>count_space UNIV" |
|
2480 |
proof (rule sums_integral) |
|
2481 |
show "\<And>i. integrable (count_space UNIV) (?f i)" |
|
2482 |
using * by (intro integrable_mult_indicator) auto |
|
2483 |
show "AE n in count_space UNIV. summable (\<lambda>i. norm (?f i n))" |
|
2484 |
using summable_finite[of "{n}" "\<lambda>i. norm (?f i n)" for n] by simp |
|
2485 |
show "summable (\<lambda>i. \<integral> n. norm (?f i n) \<partial>count_space UNIV)" |
|
2486 |
using * by (subst f') (simp add: integrable_count_space_nat_iff) |
|
2487 |
qed |
|
2488 |
also have "(\<integral> n. (\<Sum>i. ?f i n) \<partial>count_space UNIV) = (\<integral>n. f n \<partial>count_space UNIV)" |
|
2489 |
using suminf_finite[of "{n}" "\<lambda>i. ?f i n" for n] by (auto intro!: integral_cong) |
|
2490 |
also have "(\<lambda>i. \<integral>n. ?f i n \<partial>count_space UNIV) = f" |
|
2491 |
by (subst f') simp |
|
2492 |
finally show ?thesis . |
|
2493 |
qed |
|
2494 |
||
2495 |
lemma integral_count_space_nat: |
|
2496 |
fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}" |
|
2497 |
shows "integrable (count_space UNIV) f \<Longrightarrow> integral\<^sup>L (count_space UNIV) f = (\<Sum>x. f x)" |
|
2498 |
using sums_integral_count_space_nat by (rule sums_unique) |
|
2499 |
||
64008
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2500 |
lemma integrable_bij_count_space: |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2501 |
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2502 |
assumes g: "bij_betw g A B" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2503 |
shows "integrable (count_space A) (\<lambda>x. f (g x)) \<longleftrightarrow> integrable (count_space B) f" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2504 |
unfolding integrable_iff_bounded by (subst nn_integral_bij_count_space[OF g]) auto |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2505 |
|
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2506 |
lemma integral_bij_count_space: |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2507 |
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2508 |
assumes g: "bij_betw g A B" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2509 |
shows "integral\<^sup>L (count_space A) (\<lambda>x. f (g x)) = integral\<^sup>L (count_space B) f" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2510 |
using g[THEN bij_betw_imp_funcset] |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2511 |
apply (subst distr_bij_count_space[OF g, symmetric]) |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2512 |
apply (intro integral_distr[symmetric]) |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2513 |
apply auto |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2514 |
done |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2515 |
|
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2516 |
lemma has_bochner_integral_count_space_nat: |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2517 |
fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2518 |
shows "has_bochner_integral (count_space UNIV) f x \<Longrightarrow> f sums x" |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2519 |
unfolding has_bochner_integral_iff by (auto intro!: sums_integral_count_space_nat) |
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents:
63941
diff
changeset
|
2520 |
|
61808 | 2521 |
subsection \<open>Point measure\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2522 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2523 |
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
|
2524 |
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
|
2525 |
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
|
2526 |
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
|
2527 |
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
|
2528 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2529 |
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
|
2530 |
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
|
2531 |
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
|
2532 |
unfolding point_measure_def |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2533 |
apply (subst density_cong[where f'="\<lambda>x. ennreal (max 0 (f x))"]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2534 |
apply (auto split: split_max simp: ennreal_neg) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2535 |
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
|
2536 |
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
|
2537 |
done |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2538 |
|
61808 | 2539 |
subsection \<open>Lebesgue integration on @{const null_measure}\<close> |
59425 | 2540 |
|
2541 |
lemma has_bochner_integral_null_measure_iff[iff]: |
|
2542 |
"has_bochner_integral (null_measure M) f 0 \<longleftrightarrow> f \<in> borel_measurable M" |
|
2543 |
by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def] |
|
2544 |
intro!: exI[of _ "\<lambda>n x. 0"] simple_bochner_integrable.intros) |
|
2545 |
||
2546 |
lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \<longleftrightarrow> f \<in> borel_measurable M" |
|
2547 |
by (auto simp add: integrable.simps) |
|
2548 |
||
2549 |
lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0" |
|
2550 |
by (cases "integrable (null_measure M) f") |
|
2551 |
(auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq) |
|
2552 |
||
61808 | 2553 |
subsection \<open>Legacy lemmas for the real-valued Lebesgue integral\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2554 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2555 |
lemma real_lebesgue_integral_def: |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
2556 |
assumes f[measurable]: "integrable M f" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2557 |
shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+x. f x \<partial>M) - enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2558 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2559 |
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
|
2560 |
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
|
2561 |
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
|
2562 |
by (intro integral_diff integrable_max integrable_minus integrable_zero f) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2563 |
also have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = enn2real (\<integral>\<^sup>+x. ennreal (f x) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2564 |
by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2565 |
also have "integral\<^sup>L M (\<lambda>x. max 0 (- f x)) = enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2566 |
by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2567 |
finally show ?thesis . |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2568 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2569 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2570 |
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
|
2571 |
"integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and> |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2572 |
(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2573 |
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
|
2574 |
proof (safe del: notI) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2575 |
assume *: "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2576 |
have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" |
56996 | 2577 |
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
|
2578 |
also note * |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2579 |
finally show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2580 |
by simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2581 |
have "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" |
56996 | 2582 |
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
|
2583 |
also note * |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2584 |
finally show "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2585 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2586 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2587 |
assume [measurable]: "f \<in> borel_measurable M" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2588 |
assume fin: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2589 |
have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (f x) + ennreal (- f x) \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2590 |
by (intro nn_integral_cong) (auto simp: abs_real_def ennreal_neg) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2591 |
also have"\<dots> = (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) + (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M)" |
56996 | 2592 |
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
|
2593 |
also have "\<dots> < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2594 |
using fin by (auto simp: less_top) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2595 |
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
|
2596 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2597 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2598 |
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
|
2599 |
assumes "integrable M f" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2600 |
shows "f \<in> borel_measurable M" "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2601 |
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
|
2602 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2603 |
lemma integrableE: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2604 |
assumes "integrable M f" |
63941
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents:
63886
diff
changeset
|
2605 |
obtains r q where "0 \<le> r" "0 \<le> q" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2606 |
"(\<integral>\<^sup>+x. ennreal (f x)\<partial>M) = ennreal r" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2607 |
"(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M) = ennreal q" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2608 |
"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
|
2609 |
using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2610 |
by (cases rule: ennreal2_cases[of "(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M)" "(\<integral>\<^sup>+x. ennreal (f x)\<partial>M)"]) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2611 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2612 |
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
|
2613 |
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
|
2614 |
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
|
2615 |
and pos: "\<And>i. AE x in M. 0 \<le> f i x" |
61969 | 2616 |
and lim: "AE x in M. (\<lambda>i. f i x) \<longlonglongrightarrow> u x" |
2617 |
and ilim: "(\<lambda>i. integral\<^sup>L M (f i)) \<longlonglongrightarrow> x" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2618 |
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
|
2619 |
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
|
2620 |
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
|
2621 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2622 |
have nn: "AE x in M. \<forall>i. 0 \<le> f i x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2623 |
using pos unfolding AE_all_countable by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2624 |
with lim have u_nn: "AE x in M. 0 \<le> u x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2625 |
by eventually_elim (auto intro: LIMSEQ_le_const) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2626 |
have [simp]: "0 \<le> x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2627 |
by (intro LIMSEQ_le_const[OF ilim] allI exI impI integral_nonneg_AE pos) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2628 |
have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = (SUP n. (\<integral>\<^sup>+ x. ennreal (f n x) \<partial>M))" |
56996 | 2629 |
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
|
2630 |
fix i |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2631 |
from mono nn show "AE x in M. ennreal (f i x) \<le> ennreal (f (Suc i) x)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2632 |
by eventually_elim (auto simp: mono_def) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2633 |
show "(\<lambda>x. ennreal (f i x)) \<in> borel_measurable M" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2634 |
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
|
2635 |
next |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2636 |
show "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = \<integral>\<^sup>+ x. (SUP i. ennreal (f i x)) \<partial>M" |
56996 | 2637 |
apply (rule nn_integral_cong_AE) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2638 |
using lim mono nn u_nn |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2639 |
apply eventually_elim |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2640 |
apply (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2641 |
done |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2642 |
qed |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2643 |
also have "\<dots> = ennreal x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2644 |
using mono i nn unfolding nn_integral_eq_integral[OF i pos] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2645 |
by (subst LIMSEQ_unique[OF LIMSEQ_SUP]) (auto simp: mono_def integral_nonneg_AE pos intro!: integral_mono_AE ilim) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2646 |
finally have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = ennreal x" . |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2647 |
moreover have "(\<integral>\<^sup>+ x. ennreal (- u x) \<partial>M) = 0" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2648 |
using u u_nn by (subst nn_integral_0_iff_AE) (auto simp add: ennreal_neg) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2649 |
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
|
2650 |
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
|
2651 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2652 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2653 |
lemma |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2654 |
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
|
2655 |
assumes f: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)" |
61969 | 2656 |
and lim: "AE x in M. (\<lambda>i. f i x) \<longlonglongrightarrow> u x" |
2657 |
and ilim: "(\<lambda>i. integral\<^sup>L M (f i)) \<longlonglongrightarrow> x" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2658 |
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
|
2659 |
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
|
2660 |
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
|
2661 |
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
|
2662 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2663 |
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
|
2664 |
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
|
2665 |
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
|
2666 |
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
|
2667 |
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
|
2668 |
using mono by (auto simp: field_simps mono_def le_fun_def) |
61969 | 2669 |
have 4: "AE x in M. (\<lambda>i. f i x - f 0 x) \<longlonglongrightarrow> u x - f 0 x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2670 |
using lim by (auto intro!: tendsto_diff) |
61969 | 2671 |
have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x \<partial>M)) \<longlonglongrightarrow> x - integral\<^sup>L M (f 0)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2672 |
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
|
2673 |
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
|
2674 |
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
|
2675 |
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
|
2676 |
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
|
2677 |
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
|
2678 |
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
|
2679 |
by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2680 |
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
|
2681 |
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
|
2682 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2683 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2684 |
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
|
2685 |
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
|
2686 |
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
|
2687 |
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
|
2688 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2689 |
have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>x. norm (f x) \<partial>M)" |
56996 | 2690 |
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
|
2691 |
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
|
2692 |
by simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2693 |
also have "\<dots> \<longleftrightarrow> emeasure M {x\<in>space M. ennreal (norm (f x)) \<noteq> 0} = 0" |
56996 | 2694 |
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
|
2695 |
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
|
2696 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2697 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2698 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2699 |
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
|
2700 |
fixes f :: "'a \<Rightarrow> real" |
61945 | 2701 |
shows "integrable M f \<Longrightarrow> (\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> emeasure M {x\<in>space M. f x \<noteq> 0} = 0" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2702 |
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
|
2703 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2704 |
lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\<lambda>x. a)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2705 |
using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong add: less_top[symmetric]) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2706 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2707 |
lemma lebesgue_integral_const[simp]: |
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
2708 |
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
|
2709 |
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
|
2710 |
proof - |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
2711 |
{ 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
|
2712 |
then have ?thesis |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2713 |
by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) } |
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
2714 |
moreover |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
2715 |
{ 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
|
2716 |
moreover |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
2717 |
{ 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
|
2718 |
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
|
2719 |
proof qed fact |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
2720 |
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
|
2721 |
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
|
2722 |
also have "\<dots> = measure M (space M) *\<^sub>R a" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2723 |
by (simp add: less_top[symmetric]) |
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
2724 |
finally have ?thesis . } |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
2725 |
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
|
2726 |
qed |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57137
diff
changeset
|
2727 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2728 |
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
|
2729 |
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
|
2730 |
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
|
2731 |
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
|
2732 |
apply assumption |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2733 |
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
|
2734 |
apply auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2735 |
done |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2736 |
|
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2737 |
lemma (in finite_measure) integral_bounded_eq_bound_then_AE: |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2738 |
assumes "AE x in M. f x \<le> (c::real)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2739 |
"integrable M f" "(\<integral>x. f x \<partial>M) = c * measure M (space M)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2740 |
shows "AE x in M. f x = c" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2741 |
apply (rule integral_ineq_eq_0_then_AE) using assms by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
64272
diff
changeset
|
2742 |
|
59000 | 2743 |
lemma integral_indicator_finite_real: |
2744 |
fixes f :: "'a \<Rightarrow> real" |
|
2745 |
assumes [simp]: "finite A" |
|
2746 |
assumes [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M" |
|
2747 |
assumes finite: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} < \<infinity>" |
|
2748 |
shows "(\<integral>x. f x * indicator A x \<partial>M) = (\<Sum>a\<in>A. f a * measure M {a})" |
|
2749 |
proof - |
|
2750 |
have "(\<integral>x. f x * indicator A x \<partial>M) = (\<integral>x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)" |
|
2751 |
proof (intro integral_cong refl) |
|
2752 |
fix x show "f x * indicator A x = (\<Sum>a\<in>A. f a * indicator {a} x)" |
|
2753 |
by (auto split: split_indicator simp: eq_commute[of x] cong: conj_cong) |
|
2754 |
qed |
|
2755 |
also have "\<dots> = (\<Sum>a\<in>A. f a * measure M {a})" |
|
64267 | 2756 |
using finite by (subst integral_sum) (auto simp add: integrable_mult_left_iff) |
59000 | 2757 |
finally show ?thesis . |
2758 |
qed |
|
2759 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2760 |
lemma (in finite_measure) ennreal_integral_real: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2761 |
assumes [measurable]: "f \<in> borel_measurable M" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2762 |
assumes ae: "AE x in M. f x \<le> ennreal B" "0 \<le> B" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2763 |
shows "ennreal (\<integral>x. enn2real (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)" |
59000 | 2764 |
proof (subst nn_integral_eq_integral[symmetric]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2765 |
show "integrable M (\<lambda>x. enn2real (f x))" |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63627
diff
changeset
|
2766 |
using ae by (intro integrable_const_bound[where B=B]) (auto simp: enn2real_leI) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2767 |
show "(\<integral>\<^sup>+ x. ennreal (enn2real (f x)) \<partial>M) = integral\<^sup>N M f" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2768 |
using ae by (intro nn_integral_cong_AE) (auto simp: le_less_trans[OF _ ennreal_less_top]) |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63627
diff
changeset
|
2769 |
qed auto |
59000 | 2770 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2771 |
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
|
2772 |
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
|
2773 |
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
|
2774 |
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
|
2775 |
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
|
2776 |
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
|
2777 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2778 |
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
|
2779 |
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
|
2780 |
moreover |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2781 |
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
|
2782 |
proof |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2783 |
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
|
2784 |
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
|
2785 |
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
|
2786 |
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
|
2787 |
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
|
2788 |
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
|
2789 |
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
|
2790 |
moreover |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2791 |
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 | 2792 |
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
|
2793 |
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
|
2794 |
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
|
2795 |
ultimately have "emeasure M A = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2796 |
by simp |
61808 | 2797 |
with \<open>(emeasure M) A \<noteq> 0\<close> show False by auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2798 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2799 |
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
|
2800 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2801 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2802 |
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
|
2803 |
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
|
2804 |
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
|
2805 |
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
|
2806 |
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
|
2807 |
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
|
2808 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2809 |
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
|
2810 |
fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}" |
59048 | 2811 |
assumes [measurable_cong]: "sets M = sets borel" and f[measurable]: "integrable M f" |
61973 | 2812 |
shows "((\<lambda>y. \<integral> x. indicator {.. y} x *\<^sub>R f x \<partial>M) \<longlongrightarrow> \<integral> x. f x \<partial>M) at_top" |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
2813 |
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
|
2814 |
fix X :: "nat \<Rightarrow> real" assume "filterlim X at_top sequentially" |
61969 | 2815 |
show "(\<lambda>n. \<integral>x. indicator {..X n} x *\<^sub>R f x \<partial>M) \<longlonglongrightarrow> integral\<^sup>L M f" |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
2816 |
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
|
2817 |
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
|
2818 |
by (rule integrable_norm) fact |
61969 | 2819 |
show "AE x in M. (\<lambda>n. indicator {..X n} x *\<^sub>R f x) \<longlonglongrightarrow> f x" |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
2820 |
proof |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
2821 |
fix x |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2822 |
from \<open>filterlim X at_top sequentially\<close> |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
2823 |
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
|
2824 |
unfolding filterlim_at_top_ge[where c=x] by auto |
61969 | 2825 |
then show "(\<lambda>n. indicator {..X n} x *\<^sub>R f x) \<longlonglongrightarrow> f x" |
61810 | 2826 |
by (intro Lim_eventually) (auto split: split_indicator elim!: eventually_mono) |
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
2827 |
qed |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
2828 |
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
|
2829 |
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
|
2830 |
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
|
2831 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2832 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2833 |
lemma |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2834 |
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
|
2835 |
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
|
2836 |
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
|
2837 |
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
|
2838 |
assumes int: "\<And>y. integrable M (\<lambda>x. f x * indicator {.. y} x)" |
61973 | 2839 |
assumes conv: "((\<lambda>y. \<integral> x. f x * indicator {.. y} x \<partial>M) \<longlongrightarrow> x) at_top" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2840 |
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
|
2841 |
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
|
2842 |
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
|
2843 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2844 |
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
|
2845 |
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
|
2846 |
{ fix x have "eventually (\<lambda>n. f x * indicator {..real n} x = f x) sequentially" |
61942 | 2847 |
by (rule eventually_sequentiallyI[of "nat \<lceil>x\<rceil>"]) |
59587
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
nipkow
parents:
59452
diff
changeset
|
2848 |
(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
|
2849 |
from filterlim_cong[OF refl refl this] |
61969 | 2850 |
have "AE x in M. (\<lambda>i. f x * indicator {..real i} x) \<longlonglongrightarrow> f x" |
58729
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents:
57514
diff
changeset
|
2851 |
by simp |
61969 | 2852 |
have "(\<lambda>i. \<integral> x. f x * indicator {..real i} x \<partial>M) \<longlonglongrightarrow> x" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2853 |
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
|
2854 |
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
|
2855 |
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
|
2856 |
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
|
2857 |
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
|
2858 |
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
|
2859 |
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
|
2860 |
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
|
2861 |
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
|
2862 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2863 |
|
61808 | 2864 |
subsection \<open>Product measure\<close> |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2865 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2866 |
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
|
2867 |
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}" |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61169
diff
changeset
|
2868 |
assumes [measurable]: "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2869 |
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
|
2870 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2871 |
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
|
2872 |
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
|
2873 |
show ?thesis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2874 |
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
|
2875 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2876 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2877 |
lemma Collect_subset [simp]: "{x\<in>A. P x} \<subseteq> A" by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2878 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2879 |
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
|
2880 |
"(\<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
|
2881 |
{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
|
2882 |
(\<lambda>x. measure M (A x)) \<in> borel_measurable N" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2883 |
unfolding measure_def by (intro measurable_emeasure borel_measurable_enn2real) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2884 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2885 |
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
|
2886 |
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}" |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61169
diff
changeset
|
2887 |
assumes f[measurable]: "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2888 |
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
|
2889 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2890 |
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
|
2891 |
then have s: "\<And>i. simple_function (N \<Otimes>\<^sub>M M) (s i)" |
61969 | 2892 |
"\<And>x y. x \<in> space N \<Longrightarrow> y \<in> space M \<Longrightarrow> (\<lambda>i. s i (x, y)) \<longlonglongrightarrow> f x y" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2893 |
"\<And>i x y. x \<in> space N \<Longrightarrow> y \<in> space M \<Longrightarrow> norm (s i (x, y)) \<le> 2 * norm (f x y)" |
62379
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents:
62093
diff
changeset
|
2894 |
by (auto simp: space_pair_measure) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2895 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2896 |
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
|
2897 |
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
|
2898 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2899 |
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
|
2900 |
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
|
2901 |
|
63040 | 2902 |
define f' where [abs_def]: "f' i x = |
2903 |
(if integrable M (f x) then simple_bochner_integral M (\<lambda>y. s i (x, y)) else 0)" for i x |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2904 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2905 |
{ 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
|
2906 |
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
|
2907 |
(\<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
|
2908 |
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
|
2909 |
unfolding simple_bochner_integral_def |
64267 | 2910 |
by (intro sum.mono_neutral_cong_left) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2911 |
(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
|
2912 |
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
|
2913 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2914 |
show ?thesis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2915 |
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
|
2916 |
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
|
2917 |
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
|
2918 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2919 |
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
|
2920 |
{ 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
|
2921 |
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
|
2922 |
by (intro integrable_norm integrable_mult_right int_f) |
61969 | 2923 |
have "(\<lambda>i. integral\<^sup>L M (\<lambda>y. s i (x, y))) \<longlonglongrightarrow> integral\<^sup>L M (f x)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2924 |
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
|
2925 |
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
|
2926 |
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
|
2927 |
using x by simp |
61969 | 2928 |
show "AE xa in M. (\<lambda>i. s i (x, xa)) \<longlonglongrightarrow> f x xa" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2929 |
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
|
2930 |
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
|
2931 |
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
|
2932 |
qed fact |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2933 |
moreover |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2934 |
{ fix i |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2935 |
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
|
2936 |
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
|
2937 |
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
|
2938 |
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
|
2939 |
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
|
2940 |
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
|
2941 |
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
|
2942 |
(auto simp: space_pair_measure dest: finite_subset) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2943 |
have "(\<integral>\<^sup>+ y. ennreal (norm (s i (x, y))) \<partial>M) \<le> (\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M)" |
56996 | 2944 |
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
|
2945 |
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
|
2946 |
using int_2f by (simp add: integrable_iff_bounded) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2947 |
finally show "(\<integral>\<^sup>+ xa. ennreal (norm (s i (x, xa))) \<partial>M) < \<infinity>" . |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2948 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2949 |
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
|
2950 |
by (rule simple_bochner_integrable_eq_integral[symmetric]) } |
61969 | 2951 |
ultimately have "(\<lambda>i. simple_bochner_integral M (\<lambda>y. s i (x, y))) \<longlonglongrightarrow> integral\<^sup>L M (f x)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2952 |
by simp } |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
2953 |
then |
61969 | 2954 |
show "(\<lambda>i. f' i x) \<longlonglongrightarrow> integral\<^sup>L M (f x)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2955 |
unfolding f'_def |
58729
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents:
57514
diff
changeset
|
2956 |
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
|
2957 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2958 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2959 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2960 |
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
|
2961 |
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
|
2962 |
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
|
2963 |
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
|
2964 |
proof - |
61169 | 2965 |
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
|
2966 |
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
|
2967 |
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
|
2968 |
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
|
2969 |
(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
|
2970 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2971 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2972 |
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
|
2973 |
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
|
2974 |
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
|
2975 |
proof - |
61169 | 2976 |
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
|
2977 |
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
|
2978 |
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
|
2979 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2980 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2981 |
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
|
2982 |
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
|
2983 |
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
|
2984 |
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
|
2985 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2986 |
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
|
2987 |
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
|
2988 |
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
|
2989 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
2990 |
|
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
2991 |
lemma (in pair_sigma_finite) Fubini_integrable: |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
2992 |
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
2993 |
assumes f[measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
2994 |
and integ1: "integrable M1 (\<lambda>x. \<integral> y. norm (f (x, y)) \<partial>M2)" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
2995 |
and integ2: "AE x in M1. integrable M2 (\<lambda>y. f (x, y))" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
2996 |
shows "integrable (M1 \<Otimes>\<^sub>M M2) f" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
2997 |
proof (rule integrableI_bounded) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
2998 |
have "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. norm (f (x, y)) \<partial>M2) \<partial>M1)" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
2999 |
by (simp add: M2.nn_integral_fst [symmetric]) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3000 |
also have "\<dots> = (\<integral>\<^sup>+ x. \<bar>\<integral>y. norm (f (x, y)) \<partial>M2\<bar> \<partial>M1)" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3001 |
apply (intro nn_integral_cong_AE) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3002 |
using integ2 |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3003 |
proof eventually_elim |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3004 |
fix x assume "integrable M2 (\<lambda>y. f (x, y))" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3005 |
then have f: "integrable M2 (\<lambda>y. norm (f (x, y)))" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3006 |
by simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3007 |
then have "(\<integral>\<^sup>+y. ennreal (norm (f (x, y))) \<partial>M2) = ennreal (LINT y|M2. norm (f (x, y)))" |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3008 |
by (rule nn_integral_eq_integral) simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3009 |
also have "\<dots> = ennreal \<bar>LINT y|M2. norm (f (x, y))\<bar>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3010 |
using f by simp |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3011 |
finally show "(\<integral>\<^sup>+y. ennreal (norm (f (x, y))) \<partial>M2) = ennreal \<bar>LINT y|M2. norm (f (x, y))\<bar>" . |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3012 |
qed |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3013 |
also have "\<dots> < \<infinity>" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3014 |
using integ1 by (simp add: integrable_iff_bounded integral_nonneg_AE) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3015 |
finally show "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) < \<infinity>" . |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3016 |
qed fact |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
3017 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3018 |
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
|
3019 |
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
|
3020 |
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
|
3021 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3022 |
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
|
3023 |
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
|
3024 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3025 |
then have "AE x in M1. emeasure M2 (Pair x -` A) \<noteq> \<infinity>" |
56996 | 3026 |
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
|
3027 |
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
|
3028 |
using sets.sets_into_space[OF A] by (auto simp: space_pair_measure) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3029 |
ultimately show ?thesis by (auto simp: less_top) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3030 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3031 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3032 |
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
|
3033 |
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
|
3034 |
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
|
3035 |
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
|
3036 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3037 |
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 | 3038 |
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
|
3039 |
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
|
3040 |
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
|
3041 |
finally have "AE x in M1. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<noteq> \<infinity>" |
56996 | 3042 |
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
|
3043 |
(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
|
3044 |
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
|
3045 |
by eventually_elim |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3046 |
(auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]] less_top) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3047 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3048 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3049 |
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
|
3050 |
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
|
3051 |
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
|
3052 |
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
|
3053 |
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
|
3054 |
proof |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3055 |
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
|
3056 |
by (rule M2.borel_measurable_lebesgue_integral) simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3057 |
have "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) \<le> (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3058 |
using AE_integrable_fst'[OF f] by (auto intro!: nn_integral_mono_AE integral_norm_bound_ennreal) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3059 |
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 | 3060 |
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
|
3061 |
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
|
3062 |
using f unfolding integrable_iff_bounded by simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3063 |
finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) < \<infinity>" . |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3064 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3065 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3066 |
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
|
3067 |
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
|
3068 |
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
|
3069 |
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
|
3070 |
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
|
3071 |
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
|
3072 |
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
|
3073 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3074 |
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
|
3075 |
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
|
3076 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3077 |
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
|
3078 |
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
|
3079 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3080 |
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
|
3081 |
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
|
3082 |
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
|
3083 |
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
|
3084 |
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
|
3085 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3086 |
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
|
3087 |
proof (subst integral_scaleR_left) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3088 |
have "(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) = |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3089 |
(\<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
|
3090 |
using emeasure_pair_measure_finite[OF base] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3091 |
by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ennreal_measure) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3092 |
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
|
3093 |
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
|
3094 |
by (subst M2.emeasure_pair_measure_alt) |
56996 | 3095 |
(auto intro!: nn_integral_cong arg_cong[where f="emeasure M2"] simp: space_pair_measure) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3096 |
finally have *: "(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) = emeasure (M1 \<Otimes>\<^sub>M M2) A" . |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3097 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3098 |
from base * show "integrable M1 (\<lambda>x. measure M2 {y \<in> space M2. (x, y) \<in> A})" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3099 |
by (simp add: integrable_iff_bounded) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3100 |
then have "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) = |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3101 |
(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3102 |
by (rule nn_integral_eq_integral[symmetric]) simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3103 |
also note * |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3104 |
finally show "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) *\<^sub>R c = measure (M1 \<Otimes>\<^sub>M M2) A *\<^sub>R c" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3105 |
using base by (simp add: emeasure_eq_ennreal_measure) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3106 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3107 |
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
|
3108 |
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
|
3109 |
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
|
3110 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3111 |
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
|
3112 |
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
|
3113 |
by auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3114 |
have "(\<integral> x. \<integral> y. f (x, y) + g (x, y) \<partial>M2 \<partial>M1) = |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3115 |
(\<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
|
3116 |
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
|
3117 |
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
|
3118 |
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
|
3119 |
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
|
3120 |
apply simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3121 |
done |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3122 |
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
|
3123 |
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
|
3124 |
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
|
3125 |
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
|
3126 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3127 |
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
|
3128 |
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
|
3129 |
by auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3130 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3131 |
show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3132 |
proof (rule LIMSEQ_unique) |
61969 | 3133 |
show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) \<longlonglongrightarrow> integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3134 |
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
|
3135 |
show "integrable (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. 2 * norm (f x))" |
57036 | 3136 |
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
|
3137 |
qed (insert lim, auto) |
61969 | 3138 |
have "(\<lambda>i. \<integral> x. \<integral> y. s i (x, y) \<partial>M2 \<partial>M1) \<longlonglongrightarrow> \<integral> x. \<integral> y. f (x, y) \<partial>M2 \<partial>M1" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3139 |
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
|
3140 |
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
|
3141 |
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
|
3142 |
with AE_space AE_integrable_fst'[OF lim(5)] |
61969 | 3143 |
show "AE x in M1. (\<lambda>i. \<integral> y. s i (x, y) \<partial>M2) \<longlonglongrightarrow> \<integral> y. f (x, y) \<partial>M2" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3144 |
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
|
3145 |
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
|
3146 |
s: "\<forall>i. integrable M2 (\<lambda>y. s i (x, y))" and f: "integrable M2 (\<lambda>y. f (x, y))" |
61969 | 3147 |
show "(\<lambda>i. \<integral> y. s i (x, y) \<partial>M2) \<longlonglongrightarrow> \<integral> y. f (x, y) \<partial>M2" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3148 |
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
|
3149 |
show "integrable M2 (\<lambda>y. 2 * norm (f (x, y)))" |
57036 | 3150 |
using f by auto |
61969 | 3151 |
show "AE xa in M2. (\<lambda>i. s i (x, xa)) \<longlonglongrightarrow> f (x, xa)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3152 |
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
|
3153 |
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
|
3154 |
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
|
3155 |
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
|
3156 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3157 |
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
|
3158 |
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
|
3159 |
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
|
3160 |
using AE_space AE_integrable_fst'[OF lim(1), of i] AE_integrable_fst'[OF lim(5)] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3161 |
proof eventually_elim |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3162 |
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
|
3163 |
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
|
3164 |
from s have "norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral>\<^sup>+y. norm (s i (x, y)) \<partial>M2)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3165 |
by (rule integral_norm_bound_ennreal) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3166 |
also have "\<dots> \<le> (\<integral>\<^sup>+y. 2 * norm (f (x, y)) \<partial>M2)" |
56996 | 3167 |
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
|
3168 |
also have "\<dots> = (\<integral>y. 2 * norm (f (x, y)) \<partial>M2)" |
56996 | 3169 |
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
|
3170 |
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
|
3171 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3172 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3173 |
qed simp_all |
61969 | 3174 |
then show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) \<longlonglongrightarrow> \<integral> x. \<integral> y. f (x, y) \<partial>M2 \<partial>M1" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3175 |
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
|
3176 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3177 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3178 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3179 |
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
|
3180 |
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}" |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61169
diff
changeset
|
3181 |
assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3182 |
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
|
3183 |
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
|
3184 |
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
|
3185 |
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
|
3186 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3187 |
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
|
3188 |
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}" |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61169
diff
changeset
|
3189 |
assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3190 |
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
|
3191 |
and integrable_snd: "integrable M2 (\<lambda>y. \<integral>x. f x y \<partial>M1)" (is "?INT") |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61169
diff
changeset
|
3192 |
and integral_snd: "(\<integral>y. (\<integral>x. f x y \<partial>M1) \<partial>M2) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (case_prod f)" (is "?EQ") |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3193 |
proof - |
61169 | 3194 |
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
|
3195 |
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
|
3196 |
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
|
3197 |
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
|
3198 |
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
|
3199 |
show ?EQ using Q.integral_fst'[OF Q_int] |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61169
diff
changeset
|
3200 |
using integral_product_swap[of "case_prod f"] by simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3201 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3202 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3203 |
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
|
3204 |
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: {banach, second_countable_topology}" |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61169
diff
changeset
|
3205 |
assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3206 |
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
|
3207 |
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
|
3208 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3209 |
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
|
3210 |
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
|
3211 |
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
|
3212 |
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
|
3213 |
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
|
3214 |
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
|
3215 |
done |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3216 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3217 |
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
|
3218 |
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
|
3219 |
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
|
3220 |
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
|
3221 |
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
|
3222 |
proof - |
61169 | 3223 |
interpret I: finite_product_sigma_finite M I by standard fact |
3224 |
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
|
3225 |
have "finite (I \<union> J)" using fin by auto |
61169 | 3226 |
interpret IJ: finite_product_sigma_finite M "I \<union> J" by standard fact |
3227 |
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
|
3228 |
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
|
3229 |
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
|
3230 |
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
|
3231 |
by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3232 |
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
|
3233 |
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
|
3234 |
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
|
3235 |
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
|
3236 |
show ?thesis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3237 |
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
|
3238 |
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
|
3239 |
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
|
3240 |
apply simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3241 |
done |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3242 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3243 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3244 |
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
|
3245 |
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
|
3246 |
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
|
3247 |
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
|
3248 |
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
|
3249 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3250 |
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
|
3251 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3252 |
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
|
3253 |
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
|
3254 |
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
|
3255 |
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
|
3256 |
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
|
3257 |
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
|
3258 |
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
|
3259 |
show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)" |
61808 | 3260 |
using measurable_comp[OF measurable_component_update f_borel, OF x \<open>i \<notin> I\<close>] |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3261 |
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
|
3262 |
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
|
3263 |
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
|
3264 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3265 |
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
|
3266 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3267 |
|
64272 | 3268 |
lemma (in product_sigma_finite) product_integrable_prod: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3269 |
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
|
3270 |
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
|
3271 |
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
|
3272 |
proof (unfold integrable_iff_bounded, intro conjI) |
61169 | 3273 |
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
|
3274 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3275 |
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
|
3276 |
using assms by simp |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3277 |
have "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) = |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3278 |
(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. ennreal (norm (f i (x i)))) \<partial>Pi\<^sub>M I M)" |
64272 | 3279 |
by (simp add: prod_norm prod_ennreal) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3280 |
also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+ x. ennreal (norm (f i x)) \<partial>M i)" |
64272 | 3281 |
using assms by (intro product_nn_integral_prod) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3282 |
also have "\<dots> < \<infinity>" |
64272 | 3283 |
using integrable by (simp add: less_top[symmetric] ennreal_prod_eq_top integrable_iff_bounded) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3284 |
finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) < \<infinity>" . |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3285 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3286 |
|
64272 | 3287 |
lemma (in product_sigma_finite) product_integral_prod: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3288 |
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
|
3289 |
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
|
3290 |
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
|
3291 |
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
|
3292 |
case empty |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3293 |
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
|
3294 |
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
|
3295 |
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
|
3296 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3297 |
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
|
3298 |
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
|
3299 |
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
|
3300 |
integrable (Pi\<^sub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))" |
64272 | 3301 |
by (intro product_integrable_prod insert(4)) (auto intro: finite_subset) |
61169 | 3302 |
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
|
3303 |
have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))" |
64272 | 3304 |
using \<open>i \<notin> I\<close> by (auto intro!: prod.cong) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3305 |
show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3306 |
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
|
3307 |
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
|
3308 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3309 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3310 |
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
|
3311 |
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
|
3312 |
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
|
3313 |
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
|
3314 |
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
|
3315 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3316 |
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
|
3317 |
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
|
3318 |
with assms show ?thesis |
56996 | 3319 |
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
|
3320 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3321 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3322 |
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
|
3323 |
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
|
3324 |
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
|
3325 |
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
|
3326 |
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
|
3327 |
proof cases |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3328 |
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
|
3329 |
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
|
3330 |
proof induct |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3331 |
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
|
3332 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3333 |
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
|
3334 |
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
|
3335 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3336 |
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
|
3337 |
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
|
3338 |
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
|
3339 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3340 |
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
|
3341 |
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
|
3342 |
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
|
3343 |
show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3344 |
proof (intro LIMSEQ_unique) |
61969 | 3345 |
show "(\<lambda>i. integral\<^sup>L N (s i)) \<longlonglongrightarrow> integral\<^sup>L N f" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3346 |
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
|
3347 |
using lim |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3348 |
apply auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3349 |
done |
61969 | 3350 |
show "(\<lambda>i. integral\<^sup>L N (s i)) \<longlonglongrightarrow> integral\<^sup>L M f" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3351 |
unfolding lim |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3352 |
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
|
3353 |
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
|
3354 |
apply auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3355 |
done |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3356 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3357 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3358 |
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
|
3359 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3360 |
hide_const (open) simple_bochner_integral |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62390
diff
changeset
|
3361 |
hide_const (open) simple_bochner_integrable |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3362 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
diff
changeset
|
3363 |
end |