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