| author | nipkow | 
| Sun, 16 Feb 2025 17:06:09 +0100 | |
| changeset 82190 | 17e900a4a438 | 
| parent 81182 | fc5066122e68 | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Nonnegative_Lebesgue_Integration.thy | 
| 42067 | 2 | Author: Johannes Hölzl, TU München | 
| 3 | Author: Armin Heller, TU München | |
| 4 | *) | |
| 38656 | 5 | |
| 61808 | 6 | section \<open>Lebesgue Integration for Nonnegative Functions\<close> | 
| 35582 | 7 | |
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 8 | theory Nonnegative_Lebesgue_Integration | 
| 47694 | 9 | imports Measure_Space Borel_Space | 
| 35582 | 10 | begin | 
| 11 | ||
| 70136 | 12 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Approximating functions\<close> | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 13 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 14 | lemma AE_upper_bound_inf_ennreal: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 15 | fixes F G::"'a \<Rightarrow> ennreal" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 16 | assumes "\<And>e. (e::real) > 0 \<Longrightarrow> AE x in M. F x \<le> G x + e" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 17 | shows "AE x in M. F x \<le> G x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 18 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 19 | have "AE x in M. \<forall>n::nat. F x \<le> G x + ennreal (1 / Suc n)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 20 | using assms by (auto simp: AE_all_countable) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 21 | then show ?thesis | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 22 | proof (eventually_elim) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 23 | fix x assume x: "\<forall>n::nat. F x \<le> G x + ennreal (1 / Suc n)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 24 | show "F x \<le> G x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 25 | proof (rule ennreal_le_epsilon) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 26 | fix e :: real assume "0 < e" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 27 | then obtain n where n: "1 / Suc n < e" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 28 | by (blast elim: nat_approx_posE) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 29 | have "F x \<le> G x + 1 / Suc n" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 30 | using x by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 31 | also have "\<dots> \<le> G x + e" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 32 | using n by (intro add_mono ennreal_leI) auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 33 | finally show "F x \<le> G x + ennreal e" . | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 34 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 35 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 36 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 37 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 38 | lemma AE_upper_bound_inf: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 39 | fixes F G::"'a \<Rightarrow> real" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 40 | assumes "\<And>e. e > 0 \<Longrightarrow> AE x in M. F x \<le> G x + e" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 41 | shows "AE x in M. F x \<le> G x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 42 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 43 | have "AE x in M. F x \<le> G x + 1/real (n+1)" for n::nat | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 44 | by (rule assms, auto) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 45 | then have "AE x in M. \<forall>n::nat \<in> UNIV. F x \<le> G x + 1/real (n+1)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 46 | by (rule AE_ball_countable', auto) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 47 | moreover | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 48 |   {
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 49 | fix x assume i: "\<forall>n::nat \<in> UNIV. F x \<le> G x + 1/real (n+1)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 50 | have "(\<lambda>n. G x + 1/real (n+1)) \<longlonglongrightarrow> G x + 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 51 | by (rule tendsto_add, simp, rule LIMSEQ_ignore_initial_segment[OF lim_1_over_n, of 1]) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 52 | then have "F x \<le> G x" using i LIMSEQ_le_const by fastforce | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 53 | } | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 54 | ultimately show ?thesis by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 55 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 56 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 57 | lemma not_AE_zero_ennreal_E: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 58 | fixes f::"'a \<Rightarrow> ennreal" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 59 | assumes "\<not> (AE x in M. f x = 0)" and [measurable]: "f \<in> borel_measurable M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 60 | shows "\<exists>A\<in>sets M. \<exists>e::real>0. emeasure M A > 0 \<and> (\<forall>x \<in> A. f x \<ge> e)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 61 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 62 |   { assume "\<not> (\<exists>e::real>0. {x \<in> space M. f x \<ge> e} \<notin> null_sets M)"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 63 | then have "0 < e \<Longrightarrow> AE x in M. f x \<le> e" for e :: real | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 64 | by (auto simp: not_le less_imp_le dest!: AE_not_in) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 65 | then have "AE x in M. f x \<le> 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 66 | by (intro AE_upper_bound_inf_ennreal[where G="\<lambda>_. 0"]) simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 67 | then have False | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 68 | using assms by auto } | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 69 |   then obtain e::real where e: "e > 0" "{x \<in> space M. f x \<ge> e} \<notin> null_sets M" by auto
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 70 |   define A where "A = {x \<in> space M. f x \<ge> e}"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 71 | have 1 [measurable]: "A \<in> sets M" unfolding A_def by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 72 | have 2: "emeasure M A > 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 73 | using e(2) A_def \<open>A \<in> sets M\<close> by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 74 | have 3: "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> e" unfolding A_def by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 75 | show ?thesis using e(1) 1 2 3 by blast | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 76 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 77 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 78 | lemma not_AE_zero_E: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 79 | fixes f::"'a \<Rightarrow> real" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 80 | assumes "AE x in M. f x \<ge> 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 81 | "\<not>(AE x in M. f x = 0)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 82 | and [measurable]: "f \<in> borel_measurable M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 83 | shows "\<exists>A e. A \<in> sets M \<and> e>0 \<and> emeasure M A > 0 \<and> (\<forall>x \<in> A. f x \<ge> e)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 84 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 85 |   have "\<exists>e. e > 0 \<and> {x \<in> space M. f x \<ge> e} \<notin> null_sets M"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 86 | proof (rule ccontr) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 87 |     assume *: "\<not>(\<exists>e. e > 0 \<and> {x \<in> space M. f x \<ge> e} \<notin> null_sets M)"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 88 |     {
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 89 | fix e::real assume "e > 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 90 |       then have "{x \<in> space M. f x \<ge> e} \<in> null_sets M" using * by blast
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 91 |       then have "AE x in M. x \<notin> {x \<in> space M. f x \<ge> e}" using AE_not_in by blast
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 92 | then have "AE x in M. f x \<le> e" by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 93 | } | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 94 | then have "AE x in M. f x \<le> 0" by (rule AE_upper_bound_inf, auto) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 95 | then have "AE x in M. f x = 0" using assms(1) by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 96 | then show False using assms(2) by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 97 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 98 |   then obtain e where e: "e>0" "{x \<in> space M. f x \<ge> e} \<notin> null_sets M" by auto
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 99 |   define A where "A = {x \<in> space M. f x \<ge> e}"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 100 | have 1 [measurable]: "A \<in> sets M" unfolding A_def by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 101 | have 2: "emeasure M A > 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 102 | using e(2) A_def \<open>A \<in> sets M\<close> by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 103 | have 3: "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> e" unfolding A_def by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 104 | show ?thesis | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 105 | using e(1) 1 2 3 by blast | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 106 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 107 | |
| 56994 | 108 | subsection "Simple function" | 
| 35582 | 109 | |
| 61808 | 110 | text \<open> | 
| 38656 | 111 | |
| 56996 | 112 | Our simple functions are not restricted to nonnegative real numbers. Instead | 
| 38656 | 113 | they are just functions with a finite range and are measurable when singleton | 
| 114 | sets are measurable. | |
| 35582 | 115 | |
| 61808 | 116 | \<close> | 
| 38656 | 117 | |
| 70136 | 118 | definition\<^marker>\<open>tag important\<close> "simple_function M g \<longleftrightarrow> | 
| 38656 | 119 | finite (g ` space M) \<and> | 
| 120 |     (\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
 | |
| 36624 | 121 | |
| 47694 | 122 | lemma simple_functionD: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 123 | assumes "simple_function M g" | 
| 40875 | 124 | shows "finite (g ` space M)" and "g -` X \<inter> space M \<in> sets M" | 
| 40871 | 125 | proof - | 
| 126 | show "finite (g ` space M)" | |
| 127 | using assms unfolding simple_function_def by auto | |
| 40875 | 128 | have "g -` X \<inter> space M = g -` (X \<inter> g`space M) \<inter> space M" by auto | 
| 129 |   also have "\<dots> = (\<Union>x\<in>X \<inter> g`space M. g-`{x} \<inter> space M)" by auto
 | |
| 130 | finally show "g -` X \<inter> space M \<in> sets M" using assms | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 131 | by (auto simp del: UN_simps simp: simple_function_def) | 
| 40871 | 132 | qed | 
| 36624 | 133 | |
| 56949 | 134 | lemma measurable_simple_function[measurable_dest]: | 
| 135 | "simple_function M f \<Longrightarrow> f \<in> measurable M (count_space UNIV)" | |
| 136 | unfolding simple_function_def measurable_def | |
| 137 | proof safe | |
| 138 |   fix A assume "finite (f ` space M)" "\<forall>x\<in>f ` space M. f -` {x} \<inter> space M \<in> sets M"
 | |
| 139 |   then have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) \<in> sets M"
 | |
| 140 | by (intro sets.finite_UN) auto | |
| 141 |   also have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) = f -` A \<inter> space M"
 | |
| 62390 | 142 | by (auto split: if_split_asm) | 
| 56949 | 143 | finally show "f -` A \<inter> space M \<in> sets M" . | 
| 144 | qed simp | |
| 145 | ||
| 146 | lemma borel_measurable_simple_function: | |
| 147 | "simple_function M f \<Longrightarrow> f \<in> borel_measurable M" | |
| 148 | by (auto dest!: measurable_simple_function simp: measurable_def) | |
| 149 | ||
| 47694 | 150 | lemma simple_function_measurable2[intro]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 151 | assumes "simple_function M f" "simple_function M g" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 152 | shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 153 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 154 | have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 155 | by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 156 | then show ?thesis using assms[THEN simple_functionD(2)] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 157 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 158 | |
| 47694 | 159 | lemma simple_function_indicator_representation: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 160 | fixes f ::"'a \<Rightarrow> ennreal" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 161 | assumes f: "simple_function M f" and x: "x \<in> space M" | 
| 38656 | 162 |   shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
 | 
| 163 | (is "?l = ?r") | |
| 164 | proof - | |
| 38705 | 165 | have "?r = (\<Sum>y \<in> f ` space M. | 
| 38656 | 166 |     (if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
 | 
| 64267 | 167 | by (auto intro!: sum.cong) | 
| 38656 | 168 |   also have "... =  f x *  indicator (f -` {f x} \<inter> space M) x"
 | 
| 71633 | 169 | using assms by (auto dest: simple_functionD) | 
| 38656 | 170 | also have "... = f x" using x by (auto simp: indicator_def) | 
| 171 | finally show ?thesis by auto | |
| 172 | qed | |
| 36624 | 173 | |
| 47694 | 174 | lemma simple_function_notspace: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 175 | "simple_function M (\<lambda>x. h x * indicator (- space M) x::ennreal)" (is "simple_function M ?h") | 
| 35692 | 176 | proof - | 
| 38656 | 177 |   have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
 | 
| 178 | hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) | |
| 179 |   have "?h -` {0} \<inter> space M = space M" by auto
 | |
| 69661 | 180 | thus ?thesis unfolding simple_function_def by (auto simp add: image_constant_conv) | 
| 38656 | 181 | qed | 
| 182 | ||
| 47694 | 183 | lemma simple_function_cong: | 
| 38656 | 184 | assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 185 | shows "simple_function M f \<longleftrightarrow> simple_function M g" | 
| 38656 | 186 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 187 |   have "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 188 | using assms by auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 189 | with assms show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 190 | by (simp add: simple_function_def cong: image_cong) | 
| 38656 | 191 | qed | 
| 192 | ||
| 47694 | 193 | lemma simple_function_cong_algebra: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 194 | assumes "sets N = sets M" "space N = space M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 195 | shows "simple_function M f \<longleftrightarrow> simple_function N f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 196 | unfolding simple_function_def assms .. | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 197 | |
| 47694 | 198 | lemma simple_function_borel_measurable: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 199 |   fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
 | 
| 38656 | 200 | assumes "f \<in> borel_measurable M" and "finite (f ` space M)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 201 | shows "simple_function M f" | 
| 38656 | 202 | using assms unfolding simple_function_def | 
| 203 | by (auto intro: borel_measurable_vimage) | |
| 204 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 205 | lemma simple_function_iff_borel_measurable: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 206 |   fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 207 | shows "simple_function M f \<longleftrightarrow> finite (f ` space M) \<and> f \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 208 | by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 209 | |
| 56949 | 210 | lemma simple_function_eq_measurable: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 211 | "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> measurable M (count_space UNIV)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 212 | using measurable_simple_function[of M f] by (fastforce simp: simple_function_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 213 | |
| 47694 | 214 | lemma simple_function_const[intro, simp]: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 215 | "simple_function M (\<lambda>x. c)" | 
| 38656 | 216 | by (auto intro: finite_subset simp: simple_function_def) | 
| 47694 | 217 | lemma simple_function_compose[intro, simp]: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 218 | assumes "simple_function M f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 219 | shows "simple_function M (g \<circ> f)" | 
| 38656 | 220 | unfolding simple_function_def | 
| 221 | proof safe | |
| 222 | show "finite ((g \<circ> f) ` space M)" | |
| 69661 | 223 | using assms unfolding simple_function_def image_comp [symmetric] by auto | 
| 38656 | 224 | next | 
| 225 | fix x assume "x \<in> space M" | |
| 226 |   let ?G = "g -` {g (f x)} \<inter> (f`space M)"
 | |
| 227 |   have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M =
 | |
| 228 |     (\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto
 | |
| 229 |   show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M"
 | |
| 230 | using assms unfolding simple_function_def * | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 231 | by (rule_tac sets.finite_UN) auto | 
| 38656 | 232 | qed | 
| 233 | ||
| 47694 | 234 | lemma simple_function_indicator[intro, simp]: | 
| 38656 | 235 | assumes "A \<in> sets M" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 236 | shows "simple_function M (indicator A)" | 
| 35692 | 237 | proof - | 
| 38656 | 238 |   have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
 | 
| 239 | by (auto simp: indicator_def) | |
| 240 | hence "finite ?S" by (rule finite_subset) simp | |
| 241 | moreover have "- A \<inter> space M = space M - A" by auto | |
| 242 | ultimately show ?thesis unfolding simple_function_def | |
| 46905 | 243 | using assms by (auto simp: indicator_def [abs_def]) | 
| 35692 | 244 | qed | 
| 245 | ||
| 47694 | 246 | lemma simple_function_Pair[intro, simp]: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 247 | assumes "simple_function M f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 248 | assumes "simple_function M g" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 249 | shows "simple_function M (\<lambda>x. (f x, g x))" (is "simple_function M ?p") | 
| 38656 | 250 | unfolding simple_function_def | 
| 251 | proof safe | |
| 252 | show "finite (?p ` space M)" | |
| 253 | using assms unfolding simple_function_def | |
| 254 | by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto | |
| 255 | next | |
| 256 | fix x assume "x \<in> space M" | |
| 257 |   have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
 | |
| 258 |       (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
 | |
| 259 | by auto | |
| 61808 | 260 |   with \<open>x \<in> space M\<close> show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> sets M"
 | 
| 38656 | 261 | using assms unfolding simple_function_def by auto | 
| 262 | qed | |
| 35692 | 263 | |
| 47694 | 264 | lemma simple_function_compose1: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 265 | assumes "simple_function M f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 266 | shows "simple_function M (\<lambda>x. g (f x))" | 
| 38656 | 267 | using simple_function_compose[OF assms, of g] | 
| 268 | by (simp add: comp_def) | |
| 35582 | 269 | |
| 47694 | 270 | lemma simple_function_compose2: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 271 | assumes "simple_function M f" and "simple_function M g" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 272 | shows "simple_function M (\<lambda>x. h (f x) (g x))" | 
| 38656 | 273 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 274 | have "simple_function M ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))" | 
| 38656 | 275 | using assms by auto | 
| 276 | thus ?thesis by (simp_all add: comp_def) | |
| 277 | qed | |
| 35582 | 278 | |
| 67399 | 279 | lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="(+)"] | 
| 280 | and simple_function_diff[intro, simp] = simple_function_compose2[where h="(-)"] | |
| 38656 | 281 | and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67399diff
changeset | 282 | and simple_function_mult[intro, simp] = simple_function_compose2[where h="(*)"] | 
| 67399 | 283 | and simple_function_div[intro, simp] = simple_function_compose2[where h="(/)"] | 
| 38656 | 284 | and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 285 | and simple_function_max[intro, simp] = simple_function_compose2[where h=max] | 
| 38656 | 286 | |
| 64267 | 287 | lemma simple_function_sum[intro, simp]: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 288 | assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 289 | shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)" | 
| 38656 | 290 | proof cases | 
| 291 | assume "finite P" from this assms show ?thesis by induct auto | |
| 292 | qed auto | |
| 35582 | 293 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 294 | lemma simple_function_ennreal[intro, simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 295 | fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 296 | shows "simple_function M (\<lambda>x. ennreal (f x))" | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 297 | by (rule simple_function_compose1[OF sf]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 298 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 299 | lemma simple_function_real_of_nat[intro, simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 300 | fixes f g :: "'a \<Rightarrow> nat" assumes sf: "simple_function M f" | 
| 56949 | 301 | shows "simple_function M (\<lambda>x. real (f x))" | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 302 | by (rule simple_function_compose1[OF sf]) | 
| 35582 | 303 | |
| 70136 | 304 | lemma\<^marker>\<open>tag important\<close> borel_measurable_implies_simple_function_sequence: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 305 | fixes u :: "'a \<Rightarrow> ennreal" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 306 | assumes u[measurable]: "u \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 307 | shows "\<exists>f. incseq f \<and> (\<forall>i. (\<forall>x. f i x < top) \<and> simple_function M (f i)) \<and> u = (SUP i. f i)" | 
| 70136 | 308 | proof - | 
| 63040 | 309 | define f where [abs_def]: | 
| 310 | "f i x = real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i" for i x | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 311 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 312 | have [simp]: "0 \<le> f i x" for i x | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 313 | by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg enn2real_nonneg) | 
| 35582 | 314 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 315 | have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 316 | by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 317 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 318 | have "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = real_of_int \<lfloor>i * 2 ^ i\<rfloor>" for i | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 319 | by (intro arg_cong[where f=real_of_int]) simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 320 | then have [simp]: "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = i * 2 ^ i" for i | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 321 | unfolding floor_of_nat by simp | 
| 35582 | 322 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 323 | have "incseq f" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 324 | proof (intro monoI le_funI) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 325 | fix m n :: nat and x assume "m \<le> n" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 326 | moreover | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 327 |     { fix d :: nat
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 328 | have "\<lfloor>2^d::real\<rfloor> * \<lfloor>2^m * enn2real (min (of_nat m) (u x))\<rfloor> \<le> | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 329 | \<lfloor>2^d * (2^m * enn2real (min (of_nat m) (u x)))\<rfloor>" | 
| 71633 | 330 | by (rule le_mult_floor) (auto) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 331 | also have "\<dots> \<le> \<lfloor>2^d * (2^m * enn2real (min (of_nat d + of_nat m) (u x)))\<rfloor>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 332 | by (intro floor_mono mult_mono enn2real_mono min.mono) | 
| 71633 | 333 | (auto simp: min_less_iff_disj of_nat_less_top) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 334 | finally have "f m x \<le> f (m + d) x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 335 | unfolding f_def | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 336 | by (auto simp: field_simps power_add * simp del: of_int_mult) } | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 337 | ultimately show "f m x \<le> f n x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 338 | by (auto simp add: le_iff_add) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 339 | qed | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 340 | then have inc_f: "incseq (\<lambda>i. ennreal (f i x))" for x | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 341 | by (auto simp: incseq_def le_fun_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 342 | then have "incseq (\<lambda>i x. ennreal (f i x))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 343 | by (auto simp: incseq_def le_fun_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 344 | moreover | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 345 | have "simple_function M (f i)" for i | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 346 | proof (rule simple_function_borel_measurable) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 347 | have "\<lfloor>enn2real (min (of_nat i) (u x)) * 2 ^ i\<rfloor> \<le> \<lfloor>int i * 2 ^ i\<rfloor>" for x | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 348 | by (cases "u x" rule: ennreal_cases) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 349 | (auto split: split_min intro!: floor_mono) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 350 |     then have "f i ` space M \<subseteq> (\<lambda>n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}"
 | 
| 71633 | 351 | unfolding floor_of_int by (auto simp: f_def intro!: imageI) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 352 | then show "finite (f i ` space M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 353 | by (rule finite_subset) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 354 | show "f i \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 355 | unfolding f_def enn2real_def by measurable | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 356 | qed | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 357 | moreover | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 358 |   { fix x
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 359 | have "(SUP i. ennreal (f i x)) = u x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 360 | proof (cases "u x" rule: ennreal_cases) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 361 | case top then show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 362 | by (simp add: f_def inf_min[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 363 | ennreal_SUP_of_nat_eq_top) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 364 | next | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 365 | case (real r) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 366 | obtain n where "r \<le> of_nat n" using real_arch_simple by auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 367 | then have min_eq_r: "\<forall>\<^sub>F x in sequentially. min (real x) r = r" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 368 | by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 369 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 370 | have "(\<lambda>i. real_of_int \<lfloor>min (real i) r * 2^i\<rfloor> / 2^i) \<longlonglongrightarrow> r" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 371 | proof (rule tendsto_sandwich) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 372 | show "(\<lambda>n. r - (1/2)^n) \<longlonglongrightarrow> r" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 373 | by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 374 | show "\<forall>\<^sub>F n in sequentially. real_of_int \<lfloor>min (real n) r * 2 ^ n\<rfloor> / 2 ^ n \<le> r" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 375 | using min_eq_r by eventually_elim (auto simp: field_simps) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 376 | have *: "r * (2 ^ n * 2 ^ n) \<le> 2^n + 2^n * real_of_int \<lfloor>r * 2 ^ n\<rfloor>" for n | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 377 | using real_of_int_floor_ge_diff_one[of "r * 2^n", THEN mult_left_mono, of "2^n"] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 378 | by (auto simp: field_simps) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 379 | show "\<forall>\<^sub>F n in sequentially. r - (1/2)^n \<le> real_of_int \<lfloor>min (real n) r * 2 ^ n\<rfloor> / 2 ^ n" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 380 | using min_eq_r by eventually_elim (insert *, auto simp: field_simps) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 381 | qed auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 382 | then have "(\<lambda>i. ennreal (f i x)) \<longlonglongrightarrow> ennreal r" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 383 | by (simp add: real f_def ennreal_of_nat_eq_real_of_nat min_ennreal) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 384 | from LIMSEQ_unique[OF LIMSEQ_SUP[OF inc_f] this] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 385 | show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 386 | by (simp add: real) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 387 | qed } | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 388 | ultimately show ?thesis | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 389 | by (intro exI [of _ "\<lambda>i x. ennreal (f i x)"]) (auto simp add: image_comp) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 390 | qed | 
| 35582 | 391 | |
| 47694 | 392 | lemma borel_measurable_implies_simple_function_sequence': | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 393 | fixes u :: "'a \<Rightarrow> ennreal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 394 | assumes u: "u \<in> borel_measurable M" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 395 | obtains f where | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 396 | "\<And>i. simple_function M (f i)" "incseq f" "\<And>i x. f i x < top" "\<And>x. (SUP i. f i x) = u x" | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 397 | using borel_measurable_implies_simple_function_sequence [OF u] | 
| 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 398 | by (metis SUP_apply) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 399 | |
| 70136 | 400 | lemma\<^marker>\<open>tag important\<close> simple_function_induct | 
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 401 | [consumes 1, case_names cong set mult add, induct set: simple_function]: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 402 | fixes u :: "'a \<Rightarrow> ennreal" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 403 | assumes u: "simple_function M u" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 404 | assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 405 | assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 406 | assumes mult: "\<And>u c. P u \<Longrightarrow> P (\<lambda>x. c * u x)" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 407 | assumes add: "\<And>u v. P u \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 408 | shows "P u" | 
| 70136 | 409 | proof (rule cong) | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 410 |   from AE_space show "AE x in M. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x"
 | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 411 | proof eventually_elim | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 412 | fix x assume x: "x \<in> space M" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 413 | from simple_function_indicator_representation[OF u x] | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 414 |     show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x" ..
 | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 415 | qed | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 416 | next | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 417 | from u have "finite (u ` space M)" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 418 | unfolding simple_function_def by auto | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 419 |   then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
 | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 420 | proof induct | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 421 | case empty show ?case | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 422 |       using set[of "{}"] by (simp add: indicator_def[abs_def])
 | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 423 | qed (auto intro!: add mult set simple_functionD u) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 424 | next | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 425 |   show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
 | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 426 | apply (subst simple_function_cong) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 427 | apply (rule simple_function_indicator_representation[symmetric]) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 428 | apply (auto intro: u) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 429 | done | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 430 | qed fact | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 431 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 432 | lemma simple_function_induct_nn[consumes 1, case_names cong set mult add]: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 433 | fixes u :: "'a \<Rightarrow> ennreal" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 434 | assumes u: "simple_function M u" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 435 | assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 436 | assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 437 | assumes mult: "\<And>u c. simple_function M u \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 438 | assumes add: "\<And>u v. simple_function M u \<Longrightarrow> P u \<Longrightarrow> simple_function M v \<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)" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 439 | shows "P u" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 440 | proof - | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 441 | show ?thesis | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 442 | proof (rule cong) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 443 | fix x assume x: "x \<in> space M" | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 444 | from simple_function_indicator_representation[OF u x] | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 445 |     show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x" ..
 | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 446 | next | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 447 |     show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
 | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 448 | apply (subst simple_function_cong) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 449 | apply (rule simple_function_indicator_representation[symmetric]) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 450 | apply (auto intro: u) | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 451 | done | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 452 | next | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 453 | from u have "finite (u ` space M)" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 454 | unfolding simple_function_def by auto | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 455 |     then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
 | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 456 | proof induct | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 457 | case empty show ?case | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 458 |         using set[of "{}"] by (simp add: indicator_def[abs_def])
 | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 459 | next | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 460 | case (insert x S) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 461 |       { fix z have "(\<Sum>y\<in>S. y * indicator (u -` {y} \<inter> space M) z) = 0 \<or>
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 462 |           x * indicator (u -` {x} \<inter> space M) z = 0"
 | 
| 64267 | 463 | using insert by (subst sum_eq_0_iff) (auto simp: indicator_def) } | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 464 | note disj = this | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 465 | from insert show ?case | 
| 64267 | 466 | by (auto intro!: add mult set simple_functionD u simple_function_sum disj) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 467 | qed | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 468 | qed fact | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 469 | qed | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 470 | |
| 70136 | 471 | lemma\<^marker>\<open>tag important\<close> borel_measurable_induct | 
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 472 | [consumes 1, case_names cong set mult add seq, induct set: borel_measurable]: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 473 | fixes u :: "'a \<Rightarrow> ennreal" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 474 | assumes u: "u \<in> borel_measurable M" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 475 | assumes cong: "\<And>f g. f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P g \<Longrightarrow> P f" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 476 | assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 477 | assumes mult': "\<And>u c. c < top \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 478 | assumes add: "\<And>u v. u \<in> borel_measurable M\<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> v x < top) \<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)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 479 | assumes seq: "\<And>U. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. x \<in> space M \<Longrightarrow> U i x < top) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> u = (SUP i. U i) \<Longrightarrow> P (SUP i. U i)" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 480 | shows "P u" | 
| 70136 | 481 | using u | 
| 482 | proof (induct rule: borel_measurable_implies_simple_function_sequence') | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 483 | fix U assume U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i x. U i x < top" and sup: "\<And>x. (SUP i. U i x) = u x" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 484 | have u_eq: "u = (SUP i. U i)" | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 485 | using u by (auto simp add: image_comp sup) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 486 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 487 | have not_inf: "\<And>x i. x \<in> space M \<Longrightarrow> U i x < top" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 488 | using U by (auto simp: image_iff eq_commute) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 489 | |
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 490 | from U have "\<And>i. U i \<in> borel_measurable M" | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 491 | by (simp add: borel_measurable_simple_function) | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 492 | |
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 493 | show "P u" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 494 | unfolding u_eq | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 495 | proof (rule seq) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 496 | fix i show "P (U i)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 497 | using \<open>simple_function M (U i)\<close> not_inf[of _ i] | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 498 | proof (induct rule: simple_function_induct_nn) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 499 | case (mult u c) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 500 | show ?case | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 501 | proof cases | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 502 |         assume "c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0)"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 503 | with mult(1) show ?thesis | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 504 |           by (intro cong[of "\<lambda>x. c * u x" "indicator {}"] set)
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 505 | (auto dest!: borel_measurable_simple_function) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 506 | next | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 507 |         assume "\<not> (c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0))"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 508 |         then obtain x where "space M \<noteq> {}" and x: "x \<in> space M" "u x \<noteq> 0" "c \<noteq> 0"
 | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 509 | by auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 510 | with mult(3)[of x] have "c < top" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 511 | by (auto simp: ennreal_mult_less_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 512 | then have u_fin: "x' \<in> space M \<Longrightarrow> u x' < top" for x' | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 513 | using mult(3)[of x'] \<open>c \<noteq> 0\<close> by (auto simp: ennreal_mult_less_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 514 | then have "P u" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 515 | by (rule mult) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 516 | with u_fin \<open>c < top\<close> mult(1) show ?thesis | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 517 | by (intro mult') (auto dest!: borel_measurable_simple_function) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 518 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 519 | qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function) | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 520 | qed fact+ | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 521 | qed | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 522 | |
| 47694 | 523 | lemma simple_function_If_set: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 524 | assumes sf: "simple_function M f" "simple_function M g" and A: "A \<inter> space M \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 525 | shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF") | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 526 | proof - | 
| 63040 | 527 |   define F where "F x = f -` {x} \<inter> space M" for x
 | 
| 528 |   define G where "G x = g -` {x} \<inter> space M" for x
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 529 | show ?thesis unfolding simple_function_def | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 530 | proof safe | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 531 | have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 532 | from finite_subset[OF this] assms | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 533 | show "finite (?IF ` space M)" unfolding simple_function_def by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 534 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 535 | fix x assume "x \<in> space M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 536 |     then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 537 | then ((F (f x) \<inter> (A \<inter> space M)) \<union> (G (f x) - (G (f x) \<inter> (A \<inter> space M)))) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 538 | else ((F (g x) \<inter> (A \<inter> space M)) \<union> (G (g x) - (G (g x) \<inter> (A \<inter> space M)))))" | 
| 62390 | 539 | using sets.sets_into_space[OF A] by (auto split: if_split_asm simp: G_def F_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 540 | have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 541 | unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 542 |     show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
 | 
| 35582 | 543 | qed | 
| 544 | qed | |
| 545 | ||
| 47694 | 546 | lemma simple_function_If: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 547 |   assumes sf: "simple_function M f" "simple_function M g" and P: "{x\<in>space M. P x} \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 548 | shows "simple_function M (\<lambda>x. if P x then f x else g x)" | 
| 35582 | 549 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 550 |   have "{x\<in>space M. P x} = {x. P x} \<inter> space M" by auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 551 |   with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
 | 
| 38656 | 552 | qed | 
| 553 | ||
| 47694 | 554 | lemma simple_function_subalgebra: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 555 | assumes "simple_function N f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 556 | and N_subalgebra: "sets N \<subseteq> sets M" "space N = space M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 557 | shows "simple_function M f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 558 | using assms unfolding simple_function_def by auto | 
| 39092 | 559 | |
| 47694 | 560 | lemma simple_function_comp: | 
| 561 | assumes T: "T \<in> measurable M M'" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 562 | and f: "simple_function M' f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 563 | shows "simple_function M (\<lambda>x. f (T x))" | 
| 41661 | 564 | proof (intro simple_function_def[THEN iffD2] conjI ballI) | 
| 565 | have "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'" | |
| 566 | using T unfolding measurable_def by auto | |
| 567 | then show "finite ((\<lambda>x. f (T x)) ` space M)" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 568 | using f unfolding simple_function_def by (auto intro: finite_subset) | 
| 41661 | 569 | fix i assume i: "i \<in> (\<lambda>x. f (T x)) ` space M" | 
| 570 | then have "i \<in> f ` space M'" | |
| 571 | using T unfolding measurable_def by auto | |
| 572 |   then have "f -` {i} \<inter> space M' \<in> sets M'"
 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 573 | using f unfolding simple_function_def by auto | 
| 41661 | 574 |   then have "T -` (f -` {i} \<inter> space M') \<inter> space M \<in> sets M"
 | 
| 575 | using T unfolding measurable_def by auto | |
| 576 |   also have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
 | |
| 577 | using T unfolding measurable_def by auto | |
| 578 |   finally show "(\<lambda>x. f (T x)) -` {i} \<inter> space M \<in> sets M" .
 | |
| 40859 | 579 | qed | 
| 580 | ||
| 56994 | 581 | subsection "Simple integral" | 
| 38656 | 582 | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80768diff
changeset | 583 | definition\<^marker>\<open>tag important\<close> simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" (\<open>integral\<^sup>S\<close>) where
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 584 |   "integral\<^sup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 585 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 586 | syntax | 
| 81142 | 587 | "_simple_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" | 
| 588 | (\<open>(\<open>open_block notation=\<open>binder integral\<close>\<close>\<integral>\<^sup>S _. _ \<partial>_)\<close> [60,61] 110) | |
| 80768 | 589 | syntax_consts | 
| 590 | "_simple_integral" == simple_integral | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 591 | translations | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 592 | "\<integral>\<^sup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)" | 
| 35582 | 593 | |
| 47694 | 594 | lemma simple_integral_cong: | 
| 38656 | 595 | assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 596 | shows "integral\<^sup>S M f = integral\<^sup>S M g" | 
| 38656 | 597 | proof - | 
| 598 | have "f ` space M = g ` space M" | |
| 599 |     "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
 | |
| 600 | using assms by (auto intro!: image_eqI) | |
| 601 | thus ?thesis unfolding simple_integral_def by simp | |
| 602 | qed | |
| 603 | ||
| 47694 | 604 | lemma simple_integral_const[simp]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 605 | "(\<integral>\<^sup>Sx. c \<partial>M) = c * (emeasure M) (space M)" | 
| 38656 | 606 | proof (cases "space M = {}")
 | 
| 607 | case True thus ?thesis unfolding simple_integral_def by simp | |
| 608 | next | |
| 609 |   case False hence "(\<lambda>x. c) ` space M = {c}" by auto
 | |
| 610 | thus ?thesis unfolding simple_integral_def by simp | |
| 35582 | 611 | qed | 
| 612 | ||
| 47694 | 613 | lemma simple_function_partition: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 614 | assumes f: "simple_function M f" and g: "simple_function M g" | 
| 56949 | 615 | assumes sub: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow> g x = g y \<Longrightarrow> f x = f y" | 
| 616 | assumes v: "\<And>x. x \<in> space M \<Longrightarrow> f x = v (g x)" | |
| 617 |   shows "integral\<^sup>S M f = (\<Sum>y\<in>g ` space M. v y * emeasure M {x\<in>space M. g x = y})"
 | |
| 618 | (is "_ = ?r") | |
| 619 | proof - | |
| 620 | from f g have [simp]: "finite (f`space M)" "finite (g`space M)" | |
| 621 | by (auto simp: simple_function_def) | |
| 622 | from f g have [measurable]: "f \<in> measurable M (count_space UNIV)" "g \<in> measurable M (count_space UNIV)" | |
| 623 | by (auto intro: measurable_simple_function) | |
| 35582 | 624 | |
| 56949 | 625 |   { fix y assume "y \<in> space M"
 | 
| 626 |     then have "f ` space M \<inter> {i. \<exists>x\<in>space M. i = f x \<and> g y = g x} = {v (g y)}"
 | |
| 627 | by (auto cong: sub simp: v[symmetric]) } | |
| 628 | note eq = this | |
| 35582 | 629 | |
| 56949 | 630 | have "integral\<^sup>S M f = | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 631 | (\<Sum>y\<in>f`space M. y * (\<Sum>z\<in>g`space M. | 
| 56949 | 632 |       if \<exists>x\<in>space M. y = f x \<and> z = g x then emeasure M {x\<in>space M. g x = z} else 0))"
 | 
| 633 | unfolding simple_integral_def | |
| 64267 | 634 | proof (safe intro!: sum.cong ennreal_mult_left_cong) | 
| 56949 | 635 | fix y assume y: "y \<in> space M" "f y \<noteq> 0" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 636 |     have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} =
 | 
| 56949 | 637 |         {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
 | 
| 638 | by auto | |
| 639 |     have eq:"(\<Union>i\<in>{z. \<exists>x\<in>space M. f y = f x \<and> z = g x}. {x \<in> space M. g x = i}) =
 | |
| 640 |         f -` {f y} \<inter> space M"
 | |
| 641 | by (auto simp: eq_commute cong: sub rev_conj_cong) | |
| 642 | have "finite (g`space M)" by simp | |
| 643 |     then have "finite {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
 | |
| 644 | by (rule rev_finite_subset) auto | |
| 645 |     then show "emeasure M (f -` {f y} \<inter> space M) =
 | |
| 646 |       (\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then emeasure M {x \<in> space M. g x = z} else 0)"
 | |
| 64267 | 647 | apply (simp add: sum.If_cases) | 
| 648 | apply (subst sum_emeasure) | |
| 56949 | 649 | apply (auto simp: disjoint_family_on_def eq) | 
| 650 | done | |
| 38656 | 651 | qed | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 652 | also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. | 
| 56949 | 653 |       if \<exists>x\<in>space M. y = f x \<and> z = g x then y * emeasure M {x\<in>space M. g x = z} else 0))"
 | 
| 64267 | 654 | by (auto intro!: sum.cong simp: sum_distrib_left) | 
| 56949 | 655 | also have "\<dots> = ?r" | 
| 66804 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 haftmann parents: 
65680diff
changeset | 656 | by (subst sum.swap) | 
| 64267 | 657 | (auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq) | 
| 56949 | 658 | finally show "integral\<^sup>S M f = ?r" . | 
| 35582 | 659 | qed | 
| 660 | ||
| 47694 | 661 | lemma simple_integral_add[simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 662 | assumes f: "simple_function M f" and "\<And>x. 0 \<le> f x" and g: "simple_function M g" and "\<And>x. 0 \<le> g x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 663 | shows "(\<integral>\<^sup>Sx. f x + g x \<partial>M) = integral\<^sup>S M f + integral\<^sup>S M g" | 
| 35582 | 664 | proof - | 
| 56949 | 665 | have "(\<integral>\<^sup>Sx. f x + g x \<partial>M) = | 
| 666 |     (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\<in>space M. (f x, g x) = y})"
 | |
| 667 | by (intro simple_function_partition) (auto intro: f g) | |
| 668 |   also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) +
 | |
| 669 |     (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y})"
 | |
| 64267 | 670 | using assms(2,4) by (auto intro!: sum.cong distrib_right simp: sum.distrib[symmetric]) | 
| 56949 | 671 |   also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. f x \<partial>M)"
 | 
| 672 | by (intro simple_function_partition[symmetric]) (auto intro: f g) | |
| 673 |   also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. g x \<partial>M)"
 | |
| 674 | by (intro simple_function_partition[symmetric]) (auto intro: f g) | |
| 675 | finally show ?thesis . | |
| 35582 | 676 | qed | 
| 677 | ||
| 64267 | 678 | lemma simple_integral_sum[simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 679 | assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i x" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 680 | assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 681 | shows "(\<integral>\<^sup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>S M (f i))" | 
| 38656 | 682 | proof cases | 
| 683 | assume "finite P" | |
| 684 | from this assms show ?thesis | |
| 71633 | 685 | by induct (auto) | 
| 38656 | 686 | qed auto | 
| 687 | ||
| 47694 | 688 | lemma simple_integral_mult[simp]: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 689 | assumes f: "simple_function M f" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 690 | shows "(\<integral>\<^sup>Sx. c * f x \<partial>M) = c * integral\<^sup>S M f" | 
| 38656 | 691 | proof - | 
| 56949 | 692 |   have "(\<integral>\<^sup>Sx. c * f x \<partial>M) = (\<Sum>y\<in>f ` space M. (c * y) * emeasure M {x\<in>space M. f x = y})"
 | 
| 693 | using f by (intro simple_function_partition) auto | |
| 694 | also have "\<dots> = c * integral\<^sup>S M f" | |
| 695 | using f unfolding simple_integral_def | |
| 64267 | 696 | by (subst sum_distrib_left) (auto simp: mult.assoc Int_def conj_commute) | 
| 56949 | 697 | finally show ?thesis . | 
| 40871 | 698 | qed | 
| 699 | ||
| 47694 | 700 | lemma simple_integral_mono_AE: | 
| 56949 | 701 | assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g" | 
| 47694 | 702 | and mono: "AE x in M. f x \<le> g x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 703 | shows "integral\<^sup>S M f \<le> integral\<^sup>S M g" | 
| 40859 | 704 | proof - | 
| 56949 | 705 |   let ?\<mu> = "\<lambda>P. emeasure M {x\<in>space M. P x}"
 | 
| 706 | have "integral\<^sup>S M f = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * ?\<mu> (\<lambda>x. (f x, g x) = y))" | |
| 707 | using f g by (intro simple_function_partition) auto | |
| 708 | also have "\<dots> \<le> (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * ?\<mu> (\<lambda>x. (f x, g x) = y))" | |
| 64267 | 709 | proof (clarsimp intro!: sum_mono) | 
| 40859 | 710 | fix x assume "x \<in> space M" | 
| 56949 | 711 | let ?M = "?\<mu> (\<lambda>y. f y = f x \<and> g y = g x)" | 
| 712 | show "f x * ?M \<le> g x * ?M" | |
| 713 | proof cases | |
| 714 | assume "?M \<noteq> 0" | |
| 715 | then have "0 < ?M" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 716 | by (simp add: less_le) | 
| 56949 | 717 | also have "\<dots> \<le> ?\<mu> (\<lambda>y. f x \<le> g x)" | 
| 80175 
200107cdd3ac
Some new simprules – and patches for proofs
 paulson <lp15@cam.ac.uk> parents: 
79583diff
changeset | 718 | using mono by (force intro: emeasure_mono_AE) | 
| 56949 | 719 | finally have "\<not> \<not> f x \<le> g x" | 
| 720 | by (intro notI) auto | |
| 721 | then show ?thesis | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 722 | by (intro mult_right_mono) auto | 
| 56949 | 723 | qed simp | 
| 40859 | 724 | qed | 
| 56949 | 725 | also have "\<dots> = integral\<^sup>S M g" | 
| 726 | using f g by (intro simple_function_partition[symmetric]) auto | |
| 727 | finally show ?thesis . | |
| 40859 | 728 | qed | 
| 729 | ||
| 47694 | 730 | lemma simple_integral_mono: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 731 | assumes "simple_function M f" and "simple_function M g" | 
| 38656 | 732 | and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 733 | shows "integral\<^sup>S M f \<le> integral\<^sup>S M g" | 
| 41705 | 734 | using assms by (intro simple_integral_mono_AE) auto | 
| 35582 | 735 | |
| 47694 | 736 | lemma simple_integral_cong_AE: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 737 | assumes "simple_function M f" and "simple_function M g" | 
| 47694 | 738 | and "AE x in M. f x = g x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 739 | shows "integral\<^sup>S M f = integral\<^sup>S M g" | 
| 40859 | 740 | using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) | 
| 741 | ||
| 47694 | 742 | lemma simple_integral_cong': | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 743 | assumes sf: "simple_function M f" "simple_function M g" | 
| 47694 | 744 |   and mea: "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0"
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 745 | shows "integral\<^sup>S M f = integral\<^sup>S M g" | 
| 40859 | 746 | proof (intro simple_integral_cong_AE sf AE_I) | 
| 47694 | 747 |   show "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0" by fact
 | 
| 40859 | 748 |   show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
 | 
| 749 | using sf[THEN borel_measurable_simple_function] by auto | |
| 750 | qed simp | |
| 751 | ||
| 47694 | 752 | lemma simple_integral_indicator: | 
| 56949 | 753 | assumes A: "A \<in> sets M" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 754 | assumes f: "simple_function M f" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 755 | shows "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) = | 
| 56949 | 756 |     (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
 | 
| 757 | proof - | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 758 |   have eq: "(\<lambda>x. (f x, indicator A x)) ` space M \<inter> {x. snd x = 1} = (\<lambda>x. (f x, 1::ennreal))`A"
 | 
| 62390 | 759 | using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: if_split_asm) | 
| 56949 | 760 |   have eq2: "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}"
 | 
| 761 | by (auto simp: image_iff) | |
| 762 | ||
| 763 | have "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) = | |
| 764 |     (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\<in>space M. (f x, indicator A x) = y})"
 | |
| 765 | using assms by (intro simple_function_partition) auto | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 766 | also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x::ennreal))`space M. | 
| 56949 | 767 |     if snd y = 1 then fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A) else 0)"
 | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67399diff
changeset | 768 | by (auto simp: indicator_def split: if_split_asm intro!: arg_cong2[where f="(*)"] arg_cong2[where f=emeasure] sum.cong) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 769 |   also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, 1::ennreal))`A. fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A))"
 | 
| 64267 | 770 | using assms by (subst sum.If_cases) (auto intro!: simple_functionD(1) simp: eq) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 771 |   also have "\<dots> = (\<Sum>y\<in>fst`(\<lambda>x. (f x, 1::ennreal))`A. y * emeasure M (f -` {y} \<inter> space M \<inter> A))"
 | 
| 64267 | 772 | by (subst sum.reindex [of fst]) (auto simp: inj_on_def) | 
| 56949 | 773 |   also have "\<dots> = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
 | 
| 774 | using A[THEN sets.sets_into_space] | |
| 64267 | 775 | by (intro sum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2) | 
| 56949 | 776 | finally show ?thesis . | 
| 38656 | 777 | qed | 
| 35582 | 778 | |
| 47694 | 779 | lemma simple_integral_indicator_only[simp]: | 
| 38656 | 780 | assumes "A \<in> sets M" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 781 | shows "integral\<^sup>S M (indicator A) = emeasure M A" | 
| 56949 | 782 | using simple_integral_indicator[OF assms, of "\<lambda>x. 1"] sets.sets_into_space[OF assms] | 
| 62390 | 783 | by (simp_all add: image_constant_conv Int_absorb1 split: if_split_asm) | 
| 35582 | 784 | |
| 47694 | 785 | lemma simple_integral_null_set: | 
| 786 | assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets M" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 787 | shows "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = 0" | 
| 38656 | 788 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 789 | have "AE x in M. indicator N x = (0 :: ennreal)" | 
| 61808 | 790 | using \<open>N \<in> null_sets M\<close> by (auto simp: indicator_def intro!: AE_I[of _ _ N]) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 791 | then have "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^sup>Sx. 0 \<partial>M)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 792 | using assms apply (intro simple_integral_cong_AE) by auto | 
| 40859 | 793 | then show ?thesis by simp | 
| 38656 | 794 | qed | 
| 35582 | 795 | |
| 47694 | 796 | lemma simple_integral_cong_AE_mult_indicator: | 
| 797 | assumes sf: "simple_function M f" and eq: "AE x in M. x \<in> S" and "S \<in> sets M" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 798 | shows "integral\<^sup>S M f = (\<integral>\<^sup>Sx. f x * indicator S x \<partial>M)" | 
| 41705 | 799 | using assms by (intro simple_integral_cong_AE) auto | 
| 35582 | 800 | |
| 47694 | 801 | lemma simple_integral_cmult_indicator: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 802 | assumes A: "A \<in> sets M" | 
| 56949 | 803 | shows "(\<integral>\<^sup>Sx. c * indicator A x \<partial>M) = c * emeasure M A" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 804 | using simple_integral_mult[OF simple_function_indicator[OF A]] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 805 | unfolding simple_integral_indicator_only[OF A] by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 806 | |
| 56996 | 807 | lemma simple_integral_nonneg: | 
| 47694 | 808 | assumes f: "simple_function M f" and ae: "AE x in M. 0 \<le> f x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 809 | shows "0 \<le> integral\<^sup>S M f" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 810 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 811 | have "integral\<^sup>S M (\<lambda>x. 0) \<le> integral\<^sup>S M f" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 812 | using simple_integral_mono_AE[OF _ f ae] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 813 | then show ?thesis by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 814 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 815 | |
| 61808 | 816 | subsection \<open>Integral on nonnegative functions\<close> | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 817 | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80768diff
changeset | 818 | definition\<^marker>\<open>tag important\<close> nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" (\<open>integral\<^sup>N\<close>) where
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 819 |   "integral\<^sup>N M f = (SUP g \<in> {g. simple_function M g \<and> g \<le> f}. integral\<^sup>S M g)"
 | 
| 35692 | 820 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 821 | syntax | 
| 81097 | 822 | "_nn_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" (\<open>(\<open>indent=2 notation=\<open>binder integral\<close>\<close>\<integral>\<^sup>+(2 _./ _)/ \<partial>_)\<close> [60,61] 110) | 
| 80768 | 823 | syntax_consts | 
| 824 | "_nn_integral" == nn_integral | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 825 | translations | 
| 56996 | 826 | "\<integral>\<^sup>+x. f \<partial>M" == "CONST nn_integral M (\<lambda>x. f)" | 
| 40872 | 827 | |
| 56996 | 828 | lemma nn_integral_def_finite: | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 829 |   "integral\<^sup>N M f = (SUP g \<in> {g. simple_function M g \<and> g \<le> f \<and> (\<forall>x. g x < top)}. integral\<^sup>S M g)"
 | 
| 69313 | 830 | (is "_ = Sup (?A ` ?f)") | 
| 56996 | 831 | unfolding nn_integral_def | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 832 | proof (safe intro!: antisym SUP_least) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 833 | fix g assume g[measurable]: "simple_function M g" "g \<le> f" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 834 | |
| 69313 | 835 | show "integral\<^sup>S M g \<le> Sup (?A ` ?f)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 836 | proof cases | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 837 | assume ae: "AE x in M. g x \<noteq> top" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 838 |     let ?G = "{x \<in> space M. g x \<noteq> top}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 839 | have "integral\<^sup>S M g = integral\<^sup>S M (\<lambda>x. g x * indicator ?G x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 840 | proof (rule simple_integral_cong_AE) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 841 | show "AE x in M. g x = g x * indicator ?G x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 842 | using ae AE_space by eventually_elim auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 843 | qed (insert g, auto) | 
| 69313 | 844 | also have "\<dots> \<le> Sup (?A ` ?f)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 845 | using g by (intro SUP_upper) (auto simp: le_fun_def less_top split: split_indicator) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 846 | finally show ?thesis . | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 847 | next | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 848 | assume nAE: "\<not> (AE x in M. g x \<noteq> top)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 849 |     then have "emeasure M {x\<in>space M. g x = top} \<noteq> 0" (is "emeasure M ?G \<noteq> 0")
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 850 | by (subst (asm) AE_iff_measurable[OF _ refl]) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 851 | then have "top = (SUP n. (\<integral>\<^sup>Sx. of_nat n * indicator ?G x \<partial>M))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 852 | by (simp add: ennreal_SUP_of_nat_eq_top ennreal_top_eq_mult_iff SUP_mult_right_ennreal[symmetric]) | 
| 69313 | 853 | also have "\<dots> \<le> Sup (?A ` ?f)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 854 | using g | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 855 | by (safe intro!: SUP_least SUP_upper) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 856 | (auto simp: le_fun_def of_nat_less_top top_unique[symmetric] split: split_indicator | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 857 | intro: order_trans[of _ "g x" "f x" for x, OF order_trans[of _ top]]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 858 | finally show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 859 | by (simp add: top_unique del: SUP_eq_top_iff Sup_eq_top_iff) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 860 | qed | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 861 | qed (auto intro: SUP_upper) | 
| 40873 | 862 | |
| 56996 | 863 | lemma nn_integral_mono_AE: | 
| 864 | assumes ae: "AE x in M. u x \<le> v x" shows "integral\<^sup>N M u \<le> integral\<^sup>N M v" | |
| 865 | unfolding nn_integral_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 866 | proof (safe intro!: SUP_mono) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 867 | fix n assume n: "simple_function M n" "n \<le> u" | 
| 74362 | 868 | from ae[THEN AE_E] obtain N | 
| 869 |     where N: "{x \<in> space M. \<not> u x \<le> v x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
 | |
| 870 | by auto | |
| 47694 | 871 | then have ae_N: "AE x in M. x \<notin> N" by (auto intro: AE_not_in) | 
| 46731 | 872 | let ?n = "\<lambda>x. n x * indicator (space M - N) x" | 
| 47694 | 873 | have "AE x in M. n x \<le> ?n x" "simple_function M ?n" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 874 | using n N ae_N by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 875 | moreover | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 876 |   { fix x have "?n x \<le> v x"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 877 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 878 | assume x: "x \<in> space M - N" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 879 | with N have "u x \<le> v x" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 880 | with n(2)[THEN le_funD, of x] x show ?thesis | 
| 62390 | 881 | by (auto simp: max_def split: if_split_asm) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 882 | qed simp } | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 883 | then have "?n \<le> v" by (auto simp: le_funI) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 884 | moreover have "integral\<^sup>S M n \<le> integral\<^sup>S M ?n" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 885 | using ae_N N n by (auto intro!: simple_integral_mono_AE) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 886 |   ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> v}. integral\<^sup>S M n \<le> integral\<^sup>S M m"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 887 | by force | 
| 38656 | 888 | qed | 
| 889 | ||
| 56996 | 890 | lemma nn_integral_mono: | 
| 891 | "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^sup>N M u \<le> integral\<^sup>N M v" | |
| 892 | by (auto intro: nn_integral_mono_AE) | |
| 40859 | 893 | |
| 60175 | 894 | lemma mono_nn_integral: "mono F \<Longrightarrow> mono (\<lambda>x. integral\<^sup>N M (F x))" | 
| 895 | by (auto simp add: mono_def le_fun_def intro!: nn_integral_mono) | |
| 896 | ||
| 56996 | 897 | lemma nn_integral_cong_AE: | 
| 898 | "AE x in M. u x = v x \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v" | |
| 899 | by (auto simp: eq_iff intro!: nn_integral_mono_AE) | |
| 40859 | 900 | |
| 56996 | 901 | lemma nn_integral_cong: | 
| 902 | "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v" | |
| 903 | by (auto intro: nn_integral_cong_AE) | |
| 40859 | 904 | |
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 905 | lemma nn_integral_cong_simp: | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 906 | "(\<And>x. x \<in> space M =simp=> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 907 | by (auto intro: nn_integral_cong simp: simp_implies_def) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 908 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 909 | lemma incseq_nn_integral: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 910 | assumes "incseq f" shows "incseq (\<lambda>i. integral\<^sup>N M (f i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 911 | proof - | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 912 | have "\<And>i x. f i x \<le> f (Suc i) x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 913 | using assms by (auto dest!: incseq_SucD simp: le_fun_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 914 | then show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 915 | by (auto intro!: incseq_SucI nn_integral_mono) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 916 | qed | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 917 | |
| 56996 | 918 | lemma nn_integral_eq_simple_integral: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 919 | assumes f: "simple_function M f" shows "integral\<^sup>N M f = integral\<^sup>S M f" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 920 | proof - | 
| 46731 | 921 | let ?f = "\<lambda>x. f x * indicator (space M) x" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 922 | have f': "simple_function M ?f" using f by auto | 
| 56996 | 923 | have "integral\<^sup>N M ?f \<le> integral\<^sup>S M ?f" using f' | 
| 924 | by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def) | |
| 925 | moreover have "integral\<^sup>S M ?f \<le> integral\<^sup>N M ?f" | |
| 926 | unfolding nn_integral_def | |
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 927 | using f' by (auto intro!: SUP_upper) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 928 | ultimately show ?thesis | 
| 56996 | 929 | by (simp cong: nn_integral_cong simple_integral_cong) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 930 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 931 | |
| 61808 | 932 | text \<open>Beppo-Levi monotone convergence theorem\<close> | 
| 56996 | 933 | lemma nn_integral_monotone_convergence_SUP: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 934 | assumes f: "incseq f" and [measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 56996 | 935 | shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 936 | proof (rule antisym) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 937 | show "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 938 | unfolding nn_integral_def_finite[of _ "\<lambda>x. SUP i. f i x"] | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 939 | proof (safe intro!: SUP_least) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 940 | fix u assume sf_u[simp]: "simple_function M u" and | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 941 | u: "u \<le> (\<lambda>x. SUP i. f i x)" and u_range: "\<forall>x. u x < top" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 942 | note sf_u[THEN borel_measurable_simple_function, measurable] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 943 | show "integral\<^sup>S M u \<le> (SUP j. \<integral>\<^sup>+x. f j x \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 944 | proof (rule ennreal_approx_unit) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 945 | fix a :: ennreal assume "a < 1" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 946 | let ?au = "\<lambda>x. a * u x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 947 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 948 |       let ?B = "\<lambda>c i. {x\<in>space M. ?au x = c \<and> c \<le> f i x}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 949 | have "integral\<^sup>S M ?au = (\<Sum>c\<in>?au`space M. c * (SUP i. emeasure M (?B c i)))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 950 | unfolding simple_integral_def | 
| 64267 | 951 | proof (intro sum.cong ennreal_mult_left_cong refl) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 952 | fix c assume "c \<in> ?au ` space M" "c \<noteq> 0" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 953 |         { fix x' assume x': "x' \<in> space M" "?au x' = c"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 954 | with \<open>c \<noteq> 0\<close> u_range have "?au x' < 1 * u x'" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 955 | by (intro ennreal_mult_strict_right_mono \<open>a < 1\<close>) (auto simp: less_le) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 956 | also have "\<dots> \<le> (SUP i. f i x')" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 957 | using u by (auto simp: le_fun_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 958 | finally have "\<exists>i. ?au x' \<le> f i x'" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 959 | by (auto simp: less_SUP_iff intro: less_imp_le) } | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 960 |         then have *: "?au -` {c} \<inter> space M = (\<Union>i. ?B c i)"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 961 | by auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 962 |         show "emeasure M (?au -` {c} \<inter> space M) = (SUP i. emeasure M (?B c i))"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 963 | unfolding * using f | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 964 | by (intro SUP_emeasure_incseq[symmetric]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 965 | (auto simp: incseq_def le_fun_def intro: order_trans) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 966 | qed | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 967 | also have "\<dots> = (SUP i. \<Sum>c\<in>?au`space M. c * emeasure M (?B c i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 968 | unfolding SUP_mult_left_ennreal using f | 
| 64267 | 969 | by (intro ennreal_SUP_sum[symmetric]) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 970 | (auto intro!: mult_mono emeasure_mono simp: incseq_def le_fun_def intro: order_trans) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 971 | also have "\<dots> \<le> (SUP i. integral\<^sup>N M (f i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 972 | proof (intro SUP_subset_mono order_refl) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 973 | fix i | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 974 | have "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) = | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 975 |           (\<integral>\<^sup>Sx. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 976 | by (subst simple_integral_indicator) | 
| 64267 | 977 | (auto intro!: sum.cong ennreal_mult_left_cong arg_cong2[where f=emeasure]) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 978 |         also have "\<dots> = (\<integral>\<^sup>+x. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 979 | by (rule nn_integral_eq_simple_integral[symmetric]) simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 980 | also have "\<dots> \<le> (\<integral>\<^sup>+x. f i x \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 981 | by (intro nn_integral_mono) (auto split: split_indicator) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 982 | finally show "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) \<le> (\<integral>\<^sup>+x. f i x \<partial>M)" . | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 983 | qed | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 984 | finally show "a * integral\<^sup>S M u \<le> (SUP i. integral\<^sup>N M (f i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 985 | by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 986 | qed | 
| 35582 | 987 | qed | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 988 | qed (auto intro!: SUP_least SUP_upper nn_integral_mono) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 989 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 990 | lemma sup_continuous_nn_integral[order_continuous_intros]: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 991 | assumes f: "\<And>y. sup_continuous (f y)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 992 | assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 993 | shows "sup_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 994 | unfolding sup_continuous_def | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 995 | proof safe | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 996 | fix C :: "nat \<Rightarrow> 'b" assume C: "incseq C" | 
| 69313 | 997 | with sup_continuous_mono[OF f] show "(\<integral>\<^sup>+ y. f y (Sup (C ` UNIV)) \<partial>M) = (SUP i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 998 | unfolding sup_continuousD[OF f C] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 999 | by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def le_fun_def) | 
| 35582 | 1000 | qed | 
| 1001 | ||
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 1002 | theorem nn_integral_monotone_convergence_SUP_AE: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1003 | assumes f: "\<And>i. AE x in M. f i x \<le> f (Suc i) x" "\<And>i. f i \<in> borel_measurable M" | 
| 56996 | 1004 | shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))" | 
| 40859 | 1005 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1006 | from f have "AE x in M. \<forall>i. f i x \<le> f (Suc i) x" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1007 | by (simp add: AE_all_countable) | 
| 74362 | 1008 | from this[THEN AE_E] obtain N | 
| 1009 |     where N: "{x \<in> space M. \<not> (\<forall>i. f i x \<le> f (Suc i) x)} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
 | |
| 1010 | by auto | |
| 46731 | 1011 | let ?f = "\<lambda>i x. if x \<in> space M - N then f i x else 0" | 
| 47694 | 1012 | have f_eq: "AE x in M. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N]) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1013 | then have "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (SUP i. ?f i x) \<partial>M)" | 
| 56996 | 1014 | by (auto intro!: nn_integral_cong_AE) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1015 | also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. ?f i x \<partial>M))" | 
| 56996 | 1016 | proof (rule nn_integral_monotone_convergence_SUP) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1017 | show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1018 |     { fix i show "(\<lambda>x. if x \<in> space M - N then f i x else 0) \<in> borel_measurable M"
 | 
| 59000 | 1019 | using f N(3) by (intro measurable_If_set) auto } | 
| 40859 | 1020 | qed | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1021 | also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))" | 
| 69313 | 1022 | using f_eq by (force intro!: arg_cong[where f = "\<lambda>f. Sup (range f)"] nn_integral_cong_AE ext) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1023 | finally show ?thesis . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1024 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1025 | |
| 56996 | 1026 | lemma nn_integral_monotone_convergence_simple: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1027 | "incseq f \<Longrightarrow> (\<And>i. simple_function M (f i)) \<Longrightarrow> (SUP i. \<integral>\<^sup>S x. f i x \<partial>M) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)" | 
| 63092 | 1028 | using nn_integral_monotone_convergence_SUP[of f M] | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1029 | by (simp add: nn_integral_eq_simple_integral[symmetric] borel_measurable_simple_function) | 
| 40859 | 1030 | |
| 47694 | 1031 | lemma SUP_simple_integral_sequences: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1032 | assumes f: "incseq f" "\<And>i. simple_function M (f i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1033 | and g: "incseq g" "\<And>i. simple_function M (g i)" | 
| 47694 | 1034 | and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1035 | shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))" | 
| 69313 | 1036 | (is "Sup (?F ` _) = Sup (?G ` _)") | 
| 38656 | 1037 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1038 | have "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)" | 
| 56996 | 1039 | using f by (rule nn_integral_monotone_convergence_simple) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1040 | also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. g i x) \<partial>M)" | 
| 56996 | 1041 | unfolding eq[THEN nn_integral_cong_AE] .. | 
| 38656 | 1042 | also have "\<dots> = (SUP i. ?G i)" | 
| 56996 | 1043 | using g by (rule nn_integral_monotone_convergence_simple[symmetric]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1044 | finally show ?thesis by simp | 
| 38656 | 1045 | qed | 
| 1046 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1047 | lemma nn_integral_const[simp]: "(\<integral>\<^sup>+ x. c \<partial>M) = c * emeasure M (space M)" | 
| 56996 | 1048 | by (subst nn_integral_eq_simple_integral) auto | 
| 38656 | 1049 | |
| 56996 | 1050 | lemma nn_integral_linear: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1051 | assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M" | 
| 56996 | 1052 | shows "(\<integral>\<^sup>+ x. a * f x + g x \<partial>M) = a * integral\<^sup>N M f + integral\<^sup>N M g" | 
| 1053 | (is "integral\<^sup>N M ?L = _") | |
| 35582 | 1054 | proof - | 
| 74362 | 1055 | obtain u | 
| 1056 | where "\<And>i. simple_function M (u i)" "incseq u" "\<And>i x. u i x < top" "\<And>x. (SUP i. u i x) = f x" | |
| 1057 | using borel_measurable_implies_simple_function_sequence' f(1) | |
| 1058 | by auto | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1059 | note u = nn_integral_monotone_convergence_simple[OF this(2,1)] this | 
| 74362 | 1060 | |
| 1061 | obtain v where | |
| 1062 | "\<And>i. simple_function M (v i)" "incseq v" "\<And>i x. v i x < top" "\<And>x. (SUP i. v i x) = g x" | |
| 1063 | using borel_measurable_implies_simple_function_sequence' g(1) | |
| 1064 | by auto | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1065 | note v = nn_integral_monotone_convergence_simple[OF this(2,1)] this | 
| 74362 | 1066 | |
| 46731 | 1067 | let ?L' = "\<lambda>i x. a * u i x + v i x" | 
| 38656 | 1068 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1069 | have "?L \<in> borel_measurable M" using assms by auto | 
| 74362 | 1070 | from borel_measurable_implies_simple_function_sequence'[OF this] | 
| 1071 | obtain l where "\<And>i. simple_function M (l i)" "incseq l" "\<And>i x. l i x < top" "\<And>x. (SUP i. l i x) = a * f x + g x" | |
| 1072 | by auto | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1073 | note l = nn_integral_monotone_convergence_simple[OF this(2,1)] this | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1074 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1075 | have inc: "incseq (\<lambda>i. a * integral\<^sup>S M (u i))" "incseq (\<lambda>i. integral\<^sup>S M (v i))" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1076 | using u v by (auto simp: incseq_Suc_iff le_fun_def intro!: add_mono mult_left_mono simple_integral_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1077 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1078 | have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1079 | proof (rule SUP_simple_integral_sequences[OF l(3,2)]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1080 | show "incseq ?L'" "\<And>i. simple_function M (?L' i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1081 | using u v unfolding incseq_Suc_iff le_fun_def | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1082 | by (auto intro!: add_mono mult_left_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1083 |     { fix x
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1084 | have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1085 | using u(3) v(3) u(4)[of _ x] v(4)[of _ x] unfolding SUP_mult_left_ennreal | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1086 | by (auto intro!: ennreal_SUP_add simp: incseq_Suc_iff le_fun_def add_mono mult_left_mono) } | 
| 47694 | 1087 | then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1088 | unfolding l(5) using u(5) v(5) by (intro AE_I2) auto | 
| 38656 | 1089 | qed | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1090 | also have "\<dots> = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1091 | using u(2) v(2) by auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1092 | finally show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1093 | unfolding l(5)[symmetric] l(1)[symmetric] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1094 | by (simp add: ennreal_SUP_add[OF inc] v u SUP_mult_left_ennreal[symmetric]) | 
| 38656 | 1095 | qed | 
| 1096 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1097 | lemma nn_integral_cmult: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. c * f x \<partial>M) = c * integral\<^sup>N M f" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1098 | using nn_integral_linear[of f M "\<lambda>x. 0" c] by simp | 
| 38656 | 1099 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1100 | lemma nn_integral_multc: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x * c \<partial>M) = integral\<^sup>N M f * c" | 
| 63092 | 1101 | unfolding mult.commute[of _ c] nn_integral_cmult by simp | 
| 41096 | 1102 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1103 | lemma nn_integral_divide: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x / c \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M) / c" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1104 | unfolding divide_ennreal_def by (rule nn_integral_multc) | 
| 59000 | 1105 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1106 | lemma nn_integral_indicator[simp]: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A x\<partial>M) = (emeasure M) A" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1107 | by (subst nn_integral_eq_simple_integral) (auto simp: simple_integral_indicator) | 
| 38656 | 1108 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1109 | lemma nn_integral_cmult_indicator: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. c * indicator A x \<partial>M) = c * emeasure M A" | 
| 71633 | 1110 | by (subst nn_integral_eq_simple_integral) (auto) | 
| 38656 | 1111 | |
| 56996 | 1112 | lemma nn_integral_indicator': | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1113 | assumes [measurable]: "A \<inter> space M \<in> sets M" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1114 | shows "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = emeasure M (A \<inter> space M)" | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1115 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1116 | have "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = (\<integral>\<^sup>+ x. indicator (A \<inter> space M) x \<partial>M)" | 
| 56996 | 1117 | by (intro nn_integral_cong) (simp split: split_indicator) | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1118 | also have "\<dots> = emeasure M (A \<inter> space M)" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1119 | by simp | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1120 | finally show ?thesis . | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1121 | qed | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1122 | |
| 62083 | 1123 | lemma nn_integral_indicator_singleton[simp]: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1124 |   assumes [measurable]: "{y} \<in> sets M" shows "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = f y * emeasure M {y}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1125 | proof - | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1126 |   have "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = (\<integral>\<^sup>+x. f y * indicator {y} x \<partial>M)"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1127 | by (auto intro!: nn_integral_cong split: split_indicator) | 
| 62083 | 1128 | then show ?thesis | 
| 1129 | by (simp add: nn_integral_cmult) | |
| 1130 | qed | |
| 1131 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1132 | lemma nn_integral_set_ennreal: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1133 | "(\<integral>\<^sup>+x. ennreal (f x) * indicator A x \<partial>M) = (\<integral>\<^sup>+x. ennreal (f x * indicator A x) \<partial>M)" | 
| 62083 | 1134 | by (rule nn_integral_cong) (simp split: split_indicator) | 
| 1135 | ||
| 1136 | lemma nn_integral_indicator_singleton'[simp]: | |
| 1137 |   assumes [measurable]: "{y} \<in> sets M"
 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1138 |   shows "(\<integral>\<^sup>+x. ennreal (f x * indicator {y} x) \<partial>M) = f y * emeasure M {y}"
 | 
| 71633 | 1139 | by (subst nn_integral_set_ennreal[symmetric]) (simp) | 
| 62083 | 1140 | |
| 56996 | 1141 | lemma nn_integral_add: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1142 | "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x + g x \<partial>M) = integral\<^sup>N M f + integral\<^sup>N M g" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1143 | using nn_integral_linear[of f M g 1] by simp | 
| 38656 | 1144 | |
| 64267 | 1145 | lemma nn_integral_sum: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1146 | "(\<And>i. i \<in> P \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<integral>\<^sup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>N M (f i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1147 | by (induction P rule: infinite_finite_induct) (auto simp: nn_integral_add) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1148 | |
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 1149 | theorem nn_integral_suminf: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1150 | assumes f: "\<And>i. f i \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1151 | shows "(\<integral>\<^sup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>N M (f i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1152 | proof - | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1153 | have all_pos: "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 | 1154 | using assms by (auto simp: AE_all_countable) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1155 | have "(\<Sum>i. integral\<^sup>N M (f i)) = (SUP n. \<Sum>i<n. integral\<^sup>N M (f i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1156 | by (rule suminf_eq_SUP) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1157 | also have "\<dots> = (SUP n. \<integral>\<^sup>+x. (\<Sum>i<n. f i x) \<partial>M)" | 
| 64267 | 1158 | unfolding nn_integral_sum[OF f] .. | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1159 | also have "\<dots> = \<integral>\<^sup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1160 | by (intro nn_integral_monotone_convergence_SUP_AE[symmetric]) | 
| 70097 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
69861diff
changeset | 1161 | (elim AE_mp, auto simp: sum_nonneg simp del: sum.lessThan_Suc intro!: AE_I2 sum_mono2) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1162 | also have "\<dots> = \<integral>\<^sup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1163 | by (intro nn_integral_cong_AE) (auto simp: suminf_eq_SUP) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1164 | finally show ?thesis by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1165 | qed | 
| 38656 | 1166 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1167 | lemma nn_integral_bound_simple_function: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1168 | assumes bnd: "\<And>x. x \<in> space M \<Longrightarrow> f x < \<infinity>" | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1169 | assumes f[measurable]: "simple_function M f" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1170 |   assumes supp: "emeasure M {x\<in>space M. f x \<noteq> 0} < \<infinity>"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1171 | shows "nn_integral M f < \<infinity>" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1172 | proof cases | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1173 |   assume "space M = {}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1174 | then have "nn_integral M f = (\<integral>\<^sup>+x. 0 \<partial>M)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1175 | by (intro nn_integral_cong) auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1176 | then show ?thesis by simp | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1177 | next | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1178 |   assume "space M \<noteq> {}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1179 | with simple_functionD(1)[OF f] bnd have bnd: "0 \<le> Max (f`space M) \<and> Max (f`space M) < \<infinity>" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1180 | by (subst Max_less_iff) (auto simp: Max_ge_iff) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1181 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1182 |   have "nn_integral M f \<le> (\<integral>\<^sup>+x. Max (f`space M) * indicator {x\<in>space M. f x \<noteq> 0} x \<partial>M)"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1183 | proof (rule nn_integral_mono) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1184 | fix x assume "x \<in> space M" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1185 |     with f show "f x \<le> Max (f ` space M) * indicator {x \<in> space M. f x \<noteq> 0} x"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1186 | by (auto split: split_indicator intro!: Max_ge simple_functionD) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1187 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1188 | also have "\<dots> < \<infinity>" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1189 | using bnd supp by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top) | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1190 | finally show ?thesis . | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1191 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1192 | |
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 1193 | theorem nn_integral_Markov_inequality: | 
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1194 | assumes u: "(\<lambda>x. u x * indicator A x) \<in> borel_measurable M" and "A \<in> sets M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1195 |   shows "(emeasure M) ({x\<in>A. 1 \<le> c * u x}) \<le> c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
 | 
| 47694 | 1196 | (is "(emeasure M) ?A \<le> _ * ?PI") | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1197 | proof - | 
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1198 | define u' where "u' = (\<lambda>x. u x * indicator A x)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1199 | have [measurable]: "u' \<in> borel_measurable M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1200 | using u unfolding u'_def . | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1201 |   have "{x\<in>space M. c * u' x \<ge> 1} \<in> sets M"
 | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1202 | by measurable | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1203 |   also have "{x\<in>space M. c * u' x \<ge> 1} = ?A"
 | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1204 | using sets.sets_into_space[OF \<open>A \<in> sets M\<close>] by (auto simp: u'_def indicator_def) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1205 | finally have "(emeasure M) ?A = (\<integral>\<^sup>+ x. indicator ?A x \<partial>M)" | 
| 56996 | 1206 | using nn_integral_indicator by simp | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1207 | also have "\<dots> \<le> (\<integral>\<^sup>+ x. c * (u x * indicator A x) \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1208 | using u by (auto intro!: nn_integral_mono_AE simp: indicator_def) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1209 | also have "\<dots> = c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1210 | using assms by (auto intro!: nn_integral_cmult) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1211 | finally show ?thesis . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1212 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1213 | |
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1214 | lemma Chernoff_ineq_nn_integral_ge: | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1215 | assumes s: "s > 0" and [measurable]: "A \<in> sets M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1216 | assumes [measurable]: "(\<lambda>x. f x * indicator A x) \<in> borel_measurable M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1217 |   shows   "emeasure M {x\<in>A. f x \<ge> a} \<le>
 | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1218 | ennreal (exp (-s * a)) * nn_integral M (\<lambda>x. ennreal (exp (s * f x)) * indicator A x)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1219 | proof - | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1220 | define f' where "f' = (\<lambda>x. f x * indicator A x)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1221 | have [measurable]: "f' \<in> borel_measurable M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1222 | using assms(3) unfolding f'_def by assumption | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1223 | have "(\<lambda>x. ennreal (exp (s * f' x)) * indicator A x) \<in> borel_measurable M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1224 | by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1225 | also have "(\<lambda>x. ennreal (exp (s * f' x)) * indicator A x) = | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1226 | (\<lambda>x. ennreal (exp (s * f x)) * indicator A x)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1227 | by (auto simp: f'_def indicator_def fun_eq_iff) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1228 | finally have meas: "\<dots> \<in> borel_measurable M" . | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1229 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1230 |   have "{x\<in>A. f x \<ge> a} = {x\<in>A. ennreal (exp (-s * a)) * ennreal (exp (s * f x)) \<ge> 1}"
 | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1231 | using s by (auto simp: exp_minus field_simps simp flip: ennreal_mult) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1232 | also have "emeasure M \<dots> \<le> ennreal (exp (-s * a)) * | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1233 | (\<integral>\<^sup>+x. ennreal (exp (s * f x)) * indicator A x \<partial>M)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1234 | by (intro order.trans[OF nn_integral_Markov_inequality] meas) auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1235 | finally show ?thesis . | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1236 | qed | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1237 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1238 | lemma Chernoff_ineq_nn_integral_le: | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1239 | assumes s: "s > 0" and [measurable]: "A \<in> sets M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1240 | 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: 
71633diff
changeset | 1241 |   shows   "emeasure M {x\<in>A. f x \<le> a} \<le>
 | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1242 | ennreal (exp (s * a)) * nn_integral M (\<lambda>x. ennreal (exp (-s * f x)) * indicator A x)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1243 | using Chernoff_ineq_nn_integral_ge[of s A M "\<lambda>x. -f x" "-a"] assms by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1244 | |
| 56996 | 1245 | lemma nn_integral_noteq_infinite: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1246 | assumes g: "g \<in> borel_measurable M" and "integral\<^sup>N M g \<noteq> \<infinity>" | 
| 47694 | 1247 | shows "AE x in M. g x \<noteq> \<infinity>" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1248 | proof (rule ccontr) | 
| 47694 | 1249 | assume c: "\<not> (AE x in M. g x \<noteq> \<infinity>)" | 
| 1250 |   have "(emeasure M) {x\<in>space M. g x = \<infinity>} \<noteq> 0"
 | |
| 1251 | using c g by (auto simp add: AE_iff_null) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1252 |   then have "0 < (emeasure M) {x\<in>space M. g x = \<infinity>}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1253 | by (auto simp: zero_less_iff_neq_zero) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1254 |   then have "\<infinity> = \<infinity> * (emeasure M) {x\<in>space M. g x = \<infinity>}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1255 | by (auto simp: ennreal_top_eq_mult_iff) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1256 |   also have "\<dots> \<le> (\<integral>\<^sup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
 | 
| 56996 | 1257 | using g by (subst nn_integral_cmult_indicator) auto | 
| 1258 | also have "\<dots> \<le> integral\<^sup>N M g" | |
| 1259 | using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1260 | finally show False | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1261 | using \<open>integral\<^sup>N M g \<noteq> \<infinity>\<close> by (auto simp: top_unique) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1262 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1263 | |
| 56996 | 1264 | lemma nn_integral_PInf: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1265 | assumes f: "f \<in> borel_measurable M" and not_Inf: "integral\<^sup>N M f \<noteq> \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1266 |   shows "emeasure M (f -` {\<infinity>} \<inter> space M) = 0"
 | 
| 56949 | 1267 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1268 |   have "\<infinity> * emeasure M (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^sup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
 | 
| 56996 | 1269 | using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1270 | also have "\<dots> \<le> integral\<^sup>N M f" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1271 | by (auto intro!: nn_integral_mono simp: indicator_def) | 
| 56996 | 1272 |   finally have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) \<le> integral\<^sup>N M f"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1273 | by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1274 | then show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1275 | using assms by (auto simp: ennreal_top_mult top_unique split: if_split_asm) | 
| 56949 | 1276 | qed | 
| 1277 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1278 | lemma simple_integral_PInf: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1279 |   "simple_function M f \<Longrightarrow> integral\<^sup>S M f \<noteq> \<infinity> \<Longrightarrow> emeasure M (f -` {\<infinity>} \<inter> space M) = 0"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1280 | by (rule nn_integral_PInf) (auto simp: nn_integral_eq_simple_integral borel_measurable_simple_function) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1281 | |
| 56996 | 1282 | lemma nn_integral_PInf_AE: | 
| 1283 | assumes "f \<in> borel_measurable M" "integral\<^sup>N M f \<noteq> \<infinity>" shows "AE x in M. f x \<noteq> \<infinity>" | |
| 56949 | 1284 | proof (rule AE_I) | 
| 1285 |   show "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
 | |
| 56996 | 1286 | by (rule nn_integral_PInf[OF assms]) | 
| 56949 | 1287 |   show "f -` {\<infinity>} \<inter> space M \<in> sets M"
 | 
| 1288 | using assms by (auto intro: borel_measurable_vimage) | |
| 1289 | qed auto | |
| 1290 | ||
| 56996 | 1291 | lemma nn_integral_diff: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1292 | assumes f: "f \<in> borel_measurable M" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1293 | and g: "g \<in> borel_measurable M" | 
| 56996 | 1294 | and fin: "integral\<^sup>N M g \<noteq> \<infinity>" | 
| 47694 | 1295 | and mono: "AE x in M. g x \<le> f x" | 
| 56996 | 1296 | shows "(\<integral>\<^sup>+ x. f x - g x \<partial>M) = integral\<^sup>N M f - integral\<^sup>N M g" | 
| 38656 | 1297 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1298 | have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1299 | using assms by auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1300 | have "AE x in M. f x = f x - g x + g x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1301 | using diff_add_cancel_ennreal mono nn_integral_noteq_infinite[OF g fin] assms by auto | 
| 56996 | 1302 | then have **: "integral\<^sup>N M f = (\<integral>\<^sup>+x. f x - g x \<partial>M) + integral\<^sup>N M g" | 
| 1303 | unfolding nn_integral_add[OF diff g, symmetric] | |
| 1304 | by (rule nn_integral_cong_AE) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1305 | show ?thesis unfolding ** | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1306 | using fin | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1307 | by (cases rule: ennreal2_cases[of "\<integral>\<^sup>+ x. f x - g x \<partial>M" "integral\<^sup>N M g"]) auto | 
| 38656 | 1308 | qed | 
| 1309 | ||
| 56996 | 1310 | lemma nn_integral_mult_bounded_inf: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1311 | assumes f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" and c: "c \<noteq> \<infinity>" and ae: "AE x in M. g x \<le> c * f x" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1312 | shows "(\<integral>\<^sup>+x. g x \<partial>M) < \<infinity>" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1313 | proof - | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1314 | have "(\<integral>\<^sup>+x. g x \<partial>M) \<le> (\<integral>\<^sup>+x. c * f x \<partial>M)" | 
| 56996 | 1315 | by (intro nn_integral_mono_AE ae) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1316 | also have "(\<integral>\<^sup>+x. c * f x \<partial>M) < \<infinity>" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1317 | using c f by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top top_unique not_less) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1318 | finally show ?thesis . | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1319 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1320 | |
| 61808 | 1321 | text \<open>Fatou's lemma: convergence theorem on limes inferior\<close> | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1322 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1323 | lemma nn_integral_monotone_convergence_INF_AE': | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1324 | assumes f: "\<And>i. AE x in M. f (Suc i) x \<le> f i x" and [measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1325 | and *: "(\<integral>\<^sup>+ x. f 0 x \<partial>M) < \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1326 | shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1327 | proof (rule ennreal_minus_cancel) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1328 | have "integral\<^sup>N M (f 0) - (\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+x. f 0 x - (INF i. f i x) \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1329 | proof (rule nn_integral_diff[symmetric]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1330 | have "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) \<le> (\<integral>\<^sup>+ x. f 0 x \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1331 | by (intro nn_integral_mono INF_lower) simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1332 | with * show "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) \<noteq> \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1333 | by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1334 | qed (auto intro: INF_lower) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1335 | also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. f 0 x - f i x) \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1336 | by (simp add: ennreal_INF_const_minus) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1337 | also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f 0 x - f i x \<partial>M))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1338 | proof (intro nn_integral_monotone_convergence_SUP_AE) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1339 | show "AE x in M. f 0 x - f i x \<le> f 0 x - f (Suc i) x" for i | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1340 | using f[of i] by eventually_elim (auto simp: ennreal_mono_minus) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1341 | qed simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1342 | also have "\<dots> = (SUP i. nn_integral M (f 0) - (\<integral>\<^sup>+x. f i x \<partial>M))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1343 | proof (subst nn_integral_diff[symmetric]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1344 | fix i | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1345 | have dec: "AE x in M. \<forall>i. f (Suc i) x \<le> f i x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1346 | unfolding AE_all_countable using f by auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1347 | then show "AE x in M. f i x \<le> f 0 x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1348 | using dec by eventually_elim (auto intro: lift_Suc_antimono_le[of "\<lambda>i. f i x" 0 i for x]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1349 | then have "(\<integral>\<^sup>+ x. f i x \<partial>M) \<le> (\<integral>\<^sup>+ x. f 0 x \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1350 | by (rule nn_integral_mono_AE) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1351 | with * show "(\<integral>\<^sup>+ x. f i x \<partial>M) \<noteq> \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1352 | by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1353 | qed (insert f, auto simp: decseq_def le_fun_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1354 | finally show "integral\<^sup>N M (f 0) - (\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1355 | integral\<^sup>N M (f 0) - (INF i. \<integral>\<^sup>+ x. f i x \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1356 | by (simp add: ennreal_INF_const_minus) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1357 | qed (insert *, auto intro!: nn_integral_mono intro: INF_lower) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1358 | |
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 1359 | theorem nn_integral_monotone_convergence_INF_AE: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1360 | fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1361 | assumes f: "\<And>i. AE x in M. f (Suc i) x \<le> f i x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1362 | and [measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1363 | and fin: "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1364 | shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))" | 
| 38656 | 1365 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1366 |   { fix f :: "nat \<Rightarrow> ennreal" and j assume "decseq f"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1367 | then have "(INF i. f i) = (INF i. f (i + j))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1368 | apply (intro INF_eq) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1369 | apply (rule_tac x="i" in bexI) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1370 | apply (auto simp: decseq_def le_fun_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1371 | done } | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1372 | note INF_shift = this | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1373 | have mono: "AE x in M. \<forall>i. f (Suc i) x \<le> f i x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1374 | using f by (auto simp: AE_all_countable) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1375 | then have "AE x in M. (INF i. f i x) = (INF n. f (n + i) x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1376 | by eventually_elim (auto intro!: decseq_SucI INF_shift) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1377 | then have "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (INF n. f (n + i) x) \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1378 | by (rule nn_integral_cong_AE) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1379 | also have "\<dots> = (INF n. (\<integral>\<^sup>+ x. f (n + i) x \<partial>M))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1380 | by (rule nn_integral_monotone_convergence_INF_AE') (insert assms, auto) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1381 | also have "\<dots> = (INF n. (\<integral>\<^sup>+ x. f n x \<partial>M))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1382 | by (intro INF_shift[symmetric] decseq_SucI nn_integral_mono_AE f) | 
| 38656 | 1383 | finally show ?thesis . | 
| 35582 | 1384 | qed | 
| 1385 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1386 | lemma nn_integral_monotone_convergence_INF_decseq: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1387 | assumes f: "decseq f" and *: "\<And>i. f i \<in> borel_measurable M" "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1388 | shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))" | 
| 76055 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
74362diff
changeset | 1389 | using nn_integral_monotone_convergence_INF_AE[of f M i, OF _ *] f by (simp add: decseq_SucD le_funD) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1390 | |
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 1391 | theorem nn_integral_liminf: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1392 | fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1393 | assumes u: "\<And>i. u i \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1394 | shows "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1395 | proof - | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 1396 |   have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (SUP n. \<integral>\<^sup>+ x. (INF i\<in>{n..}. u i x) \<partial>M)"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1397 | unfolding liminf_SUP_INF using u | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1398 | by (intro nn_integral_monotone_convergence_SUP_AE) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1399 | (auto intro!: AE_I2 intro: INF_greatest INF_superset_mono) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1400 | also have "\<dots> \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1401 | by (auto simp: liminf_SUP_INF intro!: SUP_mono INF_greatest nn_integral_mono INF_lower) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1402 | finally show ?thesis . | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1403 | qed | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1404 | |
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 1405 | theorem nn_integral_limsup: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1406 | fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1407 | assumes [measurable]: "\<And>i. u i \<in> borel_measurable M" "w \<in> borel_measurable M" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1408 | assumes bounds: "\<And>i. AE x in M. u i x \<le> w x" and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>" | 
| 56996 | 1409 | shows "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1410 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1411 | have bnd: "AE x in M. \<forall>i. u i 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: 
56949diff
changeset | 1412 | using bounds by (auto simp: AE_all_countable) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1413 | then have "(\<integral>\<^sup>+ x. (SUP n. u n 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 | 1414 | by (auto intro!: nn_integral_mono_AE elim: eventually_mono intro: SUP_least) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 1415 |   then have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (INF n. \<integral>\<^sup>+ x. (SUP i\<in>{n..}. u i x) \<partial>M)"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1416 | unfolding limsup_INF_SUP using bnd w | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1417 | by (intro nn_integral_monotone_convergence_INF_AE') | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1418 | (auto intro!: AE_I2 intro: SUP_least SUP_subset_mono) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1419 | also have "\<dots> \<ge> limsup (\<lambda>n. integral\<^sup>N M (u n))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1420 | by (auto simp: limsup_INF_SUP intro!: INF_mono SUP_least exI nn_integral_mono SUP_upper) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1421 | finally (xtrans) show ?thesis . | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1422 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1423 | |
| 57025 | 1424 | lemma nn_integral_LIMSEQ: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1425 | assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" | 
| 61969 | 1426 | and u: "\<And>x. (\<lambda>i. f i x) \<longlonglongrightarrow> u x" | 
| 1427 | shows "(\<lambda>n. integral\<^sup>N M (f n)) \<longlonglongrightarrow> integral\<^sup>N M u" | |
| 57025 | 1428 | proof - | 
| 61969 | 1429 | have "(\<lambda>n. integral\<^sup>N M (f n)) \<longlonglongrightarrow> (SUP n. integral\<^sup>N M (f n))" | 
| 57025 | 1430 | using f by (intro LIMSEQ_SUP[of "\<lambda>n. integral\<^sup>N M (f n)"] incseq_nn_integral) | 
| 1431 | also have "(SUP n. integral\<^sup>N M (f n)) = integral\<^sup>N M (\<lambda>x. SUP n. f n x)" | |
| 1432 | using f by (intro nn_integral_monotone_convergence_SUP[symmetric]) | |
| 1433 | also have "integral\<^sup>N M (\<lambda>x. SUP n. f n x) = integral\<^sup>N M (\<lambda>x. u x)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1434 | using f by (subst LIMSEQ_SUP[THEN LIMSEQ_unique, OF _ u]) (auto simp: incseq_def le_fun_def) | 
| 57025 | 1435 | finally show ?thesis . | 
| 1436 | qed | |
| 1437 | ||
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 1438 | theorem 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: 
56949diff
changeset | 1439 | assumes [measurable]: | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1440 | "\<And>i. u i \<in> borel_measurable M" "u' \<in> borel_measurable M" "w \<in> borel_measurable M" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1441 | and bound: "\<And>j. AE x in M. 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: 
56949diff
changeset | 1442 | and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>" | 
| 61969 | 1443 | and u': "AE x in M. (\<lambda>i. u i x) \<longlonglongrightarrow> u' x" | 
| 1444 | shows "(\<lambda>i. (\<integral>\<^sup>+x. u i x \<partial>M)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. u' x \<partial>M)" | |
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1445 | proof - | 
| 56996 | 1446 | have "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)" | 
| 1447 | by (intro nn_integral_limsup[OF _ _ bound w]) auto | |
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1448 | moreover have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)" | 
| 56996 | 1449 | using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1450 | moreover have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)" | 
| 56996 | 1451 | using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) | 
| 1452 | moreover have "(\<integral>\<^sup>+x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1453 | by (intro nn_integral_liminf) auto | 
| 56996 | 1454 | moreover have "liminf (\<lambda>n. integral\<^sup>N M (u n)) \<le> limsup (\<lambda>n. integral\<^sup>N M (u n))" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1455 | by (intro Liminf_le_Limsup sequentially_bot) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1456 | ultimately show ?thesis | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1457 | by (intro Liminf_eq_Limsup) auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1458 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1459 | |
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1460 | lemma inf_continuous_nn_integral[order_continuous_intros]: | 
| 60175 | 1461 | assumes f: "\<And>y. inf_continuous (f y)" | 
| 60614 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1462 | assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M" | 
| 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1463 | assumes bnd: "\<And>x. (\<integral>\<^sup>+ y. f y x \<partial>M) \<noteq> \<infinity>" | 
| 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1464 | shows "inf_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))" | 
| 60175 | 1465 | unfolding inf_continuous_def | 
| 1466 | proof safe | |
| 60614 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1467 | fix C :: "nat \<Rightarrow> 'b" assume C: "decseq C" | 
| 69313 | 1468 | then show "(\<integral>\<^sup>+ y. f y (Inf (C ` UNIV)) \<partial>M) = (INF i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1469 | using inf_continuous_mono[OF f] bnd | 
| 76055 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
74362diff
changeset | 1470 | by (auto simp add: inf_continuousD[OF f C] fun_eq_iff monotone_def le_fun_def less_top | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1471 | intro!: nn_integral_monotone_convergence_INF_decseq) | 
| 60175 | 1472 | qed | 
| 1473 | ||
| 56996 | 1474 | lemma nn_integral_null_set: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1475 | assumes "N \<in> null_sets M" shows "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = 0" | 
| 38656 | 1476 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1477 | have "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)" | 
| 56996 | 1478 | proof (intro nn_integral_cong_AE AE_I) | 
| 40859 | 1479 |     show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N"
 | 
| 1480 | by (auto simp: indicator_def) | |
| 47694 | 1481 | show "(emeasure M) N = 0" "N \<in> sets M" | 
| 40859 | 1482 | using assms by auto | 
| 35582 | 1483 | qed | 
| 40859 | 1484 | then show ?thesis by simp | 
| 38656 | 1485 | qed | 
| 35582 | 1486 | |
| 56996 | 1487 | lemma nn_integral_0_iff: | 
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1488 | assumes u [measurable]: "u \<in> borel_measurable M" | 
| 56996 | 1489 |   shows "integral\<^sup>N M u = 0 \<longleftrightarrow> emeasure M {x\<in>space M. u x \<noteq> 0} = 0"
 | 
| 47694 | 1490 | (is "_ \<longleftrightarrow> (emeasure M) ?A = 0") | 
| 35582 | 1491 | proof - | 
| 56996 | 1492 | have u_eq: "(\<integral>\<^sup>+ x. u x * indicator ?A x \<partial>M) = integral\<^sup>N M u" | 
| 1493 | by (auto intro!: nn_integral_cong simp: indicator_def) | |
| 38656 | 1494 | show ?thesis | 
| 1495 | proof | |
| 47694 | 1496 | assume "(emeasure M) ?A = 0" | 
| 56996 | 1497 | with nn_integral_null_set[of ?A M u] u | 
| 1498 | show "integral\<^sup>N M u = 0" by (simp add: u_eq null_sets_def) | |
| 38656 | 1499 | next | 
| 56996 | 1500 | assume *: "integral\<^sup>N M u = 0" | 
| 46731 | 1501 |     let ?M = "\<lambda>n. {x \<in> space M. 1 \<le> real (n::nat) * u x}"
 | 
| 47694 | 1502 | have "0 = (SUP n. (emeasure M) (?M n \<inter> ?A))" | 
| 38656 | 1503 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1504 |       { fix n :: nat
 | 
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1505 |         have "emeasure M {x \<in> ?A. 1 \<le> of_nat n * u x} \<le>
 | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1506 | of_nat n * \<integral>\<^sup>+ x. u x * indicator ?A x \<partial>M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1507 | by (intro nn_integral_Markov_inequality) auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1508 |         also have "{x \<in> ?A. 1 \<le> of_nat n * u x} = (?M n \<inter> ?A)"
 | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1509 | by (auto simp: ennreal_of_nat_eq_real_of_nat u_eq * ) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1510 | finally have "emeasure M (?M n \<inter> ?A) \<le> 0" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1511 | by (simp add: ennreal_of_nat_eq_real_of_nat u_eq * ) | 
| 47694 | 1512 | moreover have "0 \<le> (emeasure M) (?M n \<inter> ?A)" using u by auto | 
| 1513 | ultimately have "(emeasure M) (?M n \<inter> ?A) = 0" by auto } | |
| 38656 | 1514 | thus ?thesis by simp | 
| 35582 | 1515 | qed | 
| 47694 | 1516 | also have "\<dots> = (emeasure M) (\<Union>n. ?M n \<inter> ?A)" | 
| 1517 | proof (safe intro!: SUP_emeasure_incseq) | |
| 38656 | 1518 | fix n show "?M n \<inter> ?A \<in> sets M" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 1519 | using u by (auto intro!: sets.Int) | 
| 38656 | 1520 | next | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1521 |       show "incseq (\<lambda>n. {x \<in> space M. 1 \<le> real n * u x} \<inter> {x \<in> space M. u x \<noteq> 0})"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1522 | proof (safe intro!: incseq_SucI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1523 | fix n :: nat and x | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1524 | assume *: "1 \<le> real n * u x" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1525 | also have "real n * u x \<le> real (Suc n) * u x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1526 | by (auto intro!: mult_right_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1527 | finally show "1 \<le> real (Suc n) * u x" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1528 | qed | 
| 38656 | 1529 | qed | 
| 47694 | 1530 |     also have "\<dots> = (emeasure M) {x\<in>space M. 0 < u x}"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1531 | proof (safe intro!: arg_cong[where f="(emeasure M)"]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1532 | fix x assume "0 < u x" and [simp, intro]: "x \<in> space M" | 
| 38656 | 1533 | show "x \<in> (\<Union>n. ?M n \<inter> ?A)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1534 | proof (cases "u x" rule: ennreal_cases) | 
| 61808 | 1535 | case (real r) with \<open>0 < u x\<close> have "0 < r" by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1536 | obtain j :: nat where "1 / r \<le> real j" using real_arch_simple .. | 
| 61808 | 1537 | hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using \<open>0 < r\<close> by auto | 
| 1538 | hence "1 \<le> real j * r" using real \<open>0 < r\<close> by auto | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1539 | thus ?thesis using \<open>0 < r\<close> real | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1540 | by (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_1[symmetric] ennreal_mult[symmetric] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1541 | simp del: ennreal_1) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1542 | qed (insert \<open>0 < u x\<close>, auto simp: ennreal_mult_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1543 | qed (auto simp: zero_less_iff_neq_zero) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1544 | finally show "emeasure M ?A = 0" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1545 | by (simp add: zero_less_iff_neq_zero) | 
| 35582 | 1546 | qed | 
| 1547 | qed | |
| 1548 | ||
| 56996 | 1549 | lemma nn_integral_0_iff_AE: | 
| 41705 | 1550 | assumes u: "u \<in> borel_measurable M" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1551 | shows "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x = 0)" | 
| 41705 | 1552 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1553 |   have sets: "{x\<in>space M. u x \<noteq> 0} \<in> sets M"
 | 
| 41705 | 1554 | using u by auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1555 | show "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x = 0)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1556 | using nn_integral_0_iff[of u] AE_iff_null[OF sets] u by auto | 
| 41705 | 1557 | qed | 
| 1558 | ||
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1559 | lemma AE_iff_nn_integral: | 
| 56996 | 1560 |   "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> integral\<^sup>N M (indicator {x. \<not> P x}) = 0"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1561 | by (subst nn_integral_0_iff_AE) (auto simp: indicator_def[abs_def]) | 
| 50001 
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
 hoelzl parents: 
49800diff
changeset | 1562 | |
| 59000 | 1563 | lemma nn_integral_less: | 
| 1564 | assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1565 | assumes f: "(\<integral>\<^sup>+x. f x \<partial>M) \<noteq> \<infinity>" | 
| 59000 | 1566 | assumes ord: "AE x in M. f x \<le> g x" "\<not> (AE x in M. g x \<le> f x)" | 
| 1567 | shows "(\<integral>\<^sup>+x. f x \<partial>M) < (\<integral>\<^sup>+x. g x \<partial>M)" | |
| 1568 | proof - | |
| 1569 | have "0 < (\<integral>\<^sup>+x. g x - f x \<partial>M)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1570 | proof (intro order_le_neq_trans notI) | 
| 59000 | 1571 | assume "0 = (\<integral>\<^sup>+x. g x - f x \<partial>M)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1572 | then have "AE x in M. g x - f x = 0" | 
| 59000 | 1573 | using nn_integral_0_iff_AE[of "\<lambda>x. g x - f x" M] by simp | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1574 | with ord(1) have "AE x in M. g x \<le> f x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1575 | by eventually_elim (auto simp: ennreal_minus_eq_0) | 
| 59000 | 1576 | with ord show False | 
| 1577 | by simp | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1578 | qed simp | 
| 59000 | 1579 | also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M) - (\<integral>\<^sup>+x. f x \<partial>M)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1580 | using f by (subst nn_integral_diff) (auto simp: ord) | 
| 59000 | 1581 | finally show ?thesis | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1582 | using f by (auto dest!: ennreal_minus_pos_iff[rotated] simp: less_top) | 
| 59000 | 1583 | qed | 
| 1584 | ||
| 56996 | 1585 | lemma nn_integral_subalgebra: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1586 | assumes f: "f \<in> borel_measurable N" | 
| 47694 | 1587 | and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A" | 
| 56996 | 1588 | shows "integral\<^sup>N N f = integral\<^sup>N M f" | 
| 39092 | 1589 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1590 | have [simp]: "\<And>f :: 'a \<Rightarrow> ennreal. f \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable M" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1591 | using N by (auto simp: measurable_def) | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1592 | have [simp]: "\<And>P. (AE x in N. P x) \<Longrightarrow> (AE x in M. P x)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1593 | using N by (auto simp add: eventually_ae_filter null_sets_def subset_eq) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1594 | have [simp]: "\<And>A. A \<in> sets N \<Longrightarrow> A \<in> sets M" | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1595 | using N by auto | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1596 | from f show ?thesis | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1597 | apply induct | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 1598 | apply (simp_all add: nn_integral_add nn_integral_cmult nn_integral_monotone_convergence_SUP N image_comp) | 
| 56996 | 1599 | apply (auto intro!: nn_integral_cong cong: nn_integral_cong simp: N(2)[symmetric]) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1600 | done | 
| 39092 | 1601 | qed | 
| 1602 | ||
| 56996 | 1603 | lemma nn_integral_nat_function: | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1604 | fixes f :: "'a \<Rightarrow> nat" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1605 | assumes "f \<in> measurable M (count_space UNIV)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1606 |   shows "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<Sum>t. emeasure M {x\<in>space M. t < f x})"
 | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1607 | proof - | 
| 63040 | 1608 |   define F where "F i = {x\<in>space M. i < f x}" for i
 | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1609 | with assms have [measurable]: "\<And>i. F i \<in> sets M" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1610 | by auto | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1611 | |
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1612 |   { fix x assume "x \<in> space M"
 | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1613 | have "(\<lambda>i. if i < f x then 1 else 0) sums (of_nat (f x)::real)" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1614 | using sums_If_finite[of "\<lambda>i. i < f x" "\<lambda>_. 1::real"] by simp | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1615 | then have "(\<lambda>i. ennreal (if i < f x then 1 else 0)) sums of_nat(f x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1616 | unfolding ennreal_of_nat_eq_real_of_nat | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1617 | by (subst sums_ennreal) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1618 | moreover have "\<And>i. ennreal (if i < f x then 1 else 0) = indicator (F i) x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1619 | using \<open>x \<in> space M\<close> by (simp add: one_ennreal_def F_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1620 | ultimately have "of_nat (f x) = (\<Sum>i. indicator (F i) x :: ennreal)" | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1621 | by (simp add: sums_iff) } | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1622 | then have "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)" | 
| 56996 | 1623 | by (simp cong: nn_integral_cong) | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1624 | also have "\<dots> = (\<Sum>i. emeasure M (F i))" | 
| 56996 | 1625 | by (simp add: nn_integral_suminf) | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1626 | finally show ?thesis | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1627 | by (simp add: F_def) | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1628 | qed | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1629 | |
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 1630 | theorem nn_integral_lfp: | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1631 | assumes sets[simp]: "\<And>s. sets (M s) = sets N" | 
| 60175 | 1632 | assumes f: "sup_continuous f" | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1633 | assumes g: "sup_continuous g" | 
| 60175 | 1634 | assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N" | 
| 1635 | assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s" | |
| 1636 | shows "(\<integral>\<^sup>+\<omega>. lfp f \<omega> \<partial>M s) = lfp g s" | |
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1637 | proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f and P="\<lambda>f. f \<in> borel_measurable N", symmetric]) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1638 | fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ennreal" assume "incseq C" "\<And>i. C i \<in> borel_measurable N" | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1639 | then show "(\<lambda>s. \<integral>\<^sup>+x. (SUP i. C i) x \<partial>M s) = (SUP i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1640 | unfolding SUP_apply[abs_def] | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1641 | by (subst nn_integral_monotone_convergence_SUP) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1642 | (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1643 | qed (auto simp add: step le_fun_def SUP_apply[abs_def] bot_fun_def bot_ennreal intro!: meas f g) | 
| 60175 | 1644 | |
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 1645 | theorem nn_integral_gfp: | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1646 | assumes sets[simp]: "\<And>s. sets (M s) = sets N" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1647 | assumes f: "inf_continuous f" and g: "inf_continuous g" | 
| 60175 | 1648 | assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N" | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1649 | assumes bound: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^sup>+x. f F x \<partial>M s) < \<infinity>" | 
| 60175 | 1650 | assumes non_zero: "\<And>s. emeasure (M s) (space (M s)) \<noteq> 0" | 
| 1651 | assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s" | |
| 1652 | shows "(\<integral>\<^sup>+\<omega>. gfp f \<omega> \<partial>M s) = gfp g s" | |
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1653 | proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1654 | and P="\<lambda>F. F \<in> borel_measurable N \<and> (\<forall>s. (\<integral>\<^sup>+x. F x \<partial>M s) < \<infinity>)", symmetric]) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1655 | fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ennreal" assume "decseq C" "\<And>i. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)" | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1656 | then show "(\<lambda>s. \<integral>\<^sup>+x. (INF i. C i) x \<partial>M s) = (INF i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1657 | unfolding INF_apply[abs_def] | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1658 | by (subst nn_integral_monotone_convergence_INF_decseq) | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1659 | (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1660 | next | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1661 | show "\<And>x. g x \<le> (\<lambda>s. integral\<^sup>N (M s) (f top))" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1662 | by (subst step) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1663 | (auto simp add: top_fun_def less_le non_zero le_fun_def ennreal_top_mult | 
| 63566 | 1664 | cong del: if_weak_cong intro!: monoD[OF inf_continuous_mono[OF g], THEN le_funD]) | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1665 | next | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1666 | fix C assume "\<And>i::nat. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)" "decseq C" | 
| 69313 | 1667 | with bound show "Inf (C ` UNIV) \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (Inf (C ` UNIV)) < \<infinity>)" | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1668 | unfolding INF_apply[abs_def] | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1669 | by (subst nn_integral_monotone_convergence_INF_decseq) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1670 | (auto simp: INF_less_iff cong: measurable_cong_sets intro!: borel_measurable_INF) | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1671 | next | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1672 | show "\<And>x. x \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) x < \<infinity>) \<Longrightarrow> | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1673 | (\<lambda>s. integral\<^sup>N (M s) (f x)) = g (\<lambda>s. integral\<^sup>N (M s) x)" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1674 | by (subst step) auto | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1675 | qed (insert bound, auto simp add: le_fun_def INF_apply[abs_def] top_fun_def intro!: meas f g) | 
| 60175 | 1676 | |
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1677 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1678 | text \<open>Cauchy--Schwarz inequality for \<^const>\<open>nn_integral\<close>\<close> | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1679 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1680 | lemma sum_of_squares_ge_ennreal: | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1681 | fixes a b :: ennreal | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1682 | shows "2 * a * b \<le> a\<^sup>2 + b\<^sup>2" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1683 | proof (cases a; cases b) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1684 | fix x y | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1685 | assume xy: "x \<ge> 0" "y \<ge> 0" and [simp]: "a = ennreal x" "b = ennreal y" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1686 | have "0 \<le> (x - y)\<^sup>2" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1687 | by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1688 | also have "\<dots> = x\<^sup>2 + y\<^sup>2 - 2 * x * y" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1689 | by (simp add: algebra_simps power2_eq_square) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1690 | finally have "2 * x * y \<le> x\<^sup>2 + y\<^sup>2" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1691 | by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1692 | hence "ennreal (2 * x * y) \<le> ennreal (x\<^sup>2 + y\<^sup>2)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1693 | by (intro ennreal_leI) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1694 | thus ?thesis using xy | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1695 | by (simp add: ennreal_mult ennreal_power) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1696 | qed auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1697 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1698 | lemma Cauchy_Schwarz_nn_integral: | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1699 | assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1700 | shows "(\<integral>\<^sup>+x. f x * g x \<partial>M)\<^sup>2 \<le> (\<integral>\<^sup>+x. f x ^ 2 \<partial>M) * (\<integral>\<^sup>+x. g x ^ 2 \<partial>M)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1701 | proof (cases "(\<integral>\<^sup>+x. f x * g x \<partial>M) = 0") | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1702 | case False | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1703 | define F where "F = nn_integral M (\<lambda>x. f x ^ 2)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1704 | define G where "G = nn_integral M (\<lambda>x. g x ^ 2)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1705 | from False have "\<not>(AE x in M. f x = 0 \<or> g x = 0)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1706 | by (auto simp: nn_integral_0_iff_AE) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1707 | hence "\<not>(AE x in M. f x = 0)" and "\<not>(AE x in M. g x = 0)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1708 | by (auto intro: AE_disjI1 AE_disjI2) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1709 | hence nz: "F \<noteq> 0" "G \<noteq> 0" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1710 | by (auto simp: nn_integral_0_iff_AE F_def G_def) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1711 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1712 | show ?thesis | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1713 | proof (cases "F = \<infinity> \<or> G = \<infinity>") | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1714 | case True | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1715 | thus ?thesis using nz | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1716 | by (auto simp: F_def G_def) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1717 | next | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1718 | case False | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1719 | define F' where "F' = ennreal (sqrt (enn2real F))" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1720 | define G' where "G' = ennreal (sqrt (enn2real G))" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1721 | from False have fin: "F < top" "G < top" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1722 | by (simp_all add: top.not_eq_extremum) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1723 | have F'_sqr: "F'\<^sup>2 = F" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1724 | using False by (cases F) (auto simp: F'_def ennreal_power) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1725 | have G'_sqr: "G'\<^sup>2 = G" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1726 | using False by (cases G) (auto simp: G'_def ennreal_power) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1727 | have nz': "F' \<noteq> 0" "G' \<noteq> 0" and fin': "F' \<noteq> \<infinity>" "G' \<noteq> \<infinity>" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1728 | using F'_sqr G'_sqr nz fin by auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1729 | from fin' have fin'': "F' < top" "G' < top" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1730 | by (auto simp: top.not_eq_extremum) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1731 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1732 | have "2 * (F' / F') * (G' / G') * (\<integral>\<^sup>+x. f x * g x \<partial>M) = | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1733 | F' * G' * (\<integral>\<^sup>+x. 2 * (f x / F') * (g x / G') \<partial>M)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1734 | using nz' fin'' | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1735 | by (simp add: divide_ennreal_def algebra_simps ennreal_inverse_mult flip: nn_integral_cmult) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1736 | also have "F'/ F' = 1" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1737 | using nz' fin'' by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1738 | also have "G'/ G' = 1" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1739 | using nz' fin'' by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1740 | also have "2 * 1 * 1 = (2 :: ennreal)" by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1741 | also have "F' * G' * (\<integral>\<^sup>+ x. 2 * (f x / F') * (g x / G') \<partial>M) \<le> | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1742 | F' * G' * (\<integral>\<^sup>+x. (f x / F')\<^sup>2 + (g x / G')\<^sup>2 \<partial>M)" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1743 | by (intro mult_left_mono nn_integral_mono sum_of_squares_ge_ennreal) auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1744 | also have "\<dots> = F' * G' * (F / F'\<^sup>2 + G / G'\<^sup>2)" using nz | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1745 | by (auto simp: nn_integral_add algebra_simps nn_integral_divide F_def G_def) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1746 | also have "F / F'\<^sup>2 = 1" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1747 | using nz F'_sqr fin by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1748 | also have "G / G'\<^sup>2 = 1" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1749 | using nz G'_sqr fin by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1750 | also have "F' * G' * (1 + 1) = 2 * (F' * G')" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1751 | by (simp add: mult_ac) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1752 | finally have "(\<integral>\<^sup>+x. f x * g x \<partial>M) \<le> F' * G'" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1753 | by (subst (asm) ennreal_mult_le_mult_iff) auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1754 | hence "(\<integral>\<^sup>+x. f x * g x \<partial>M)\<^sup>2 \<le> (F' * G')\<^sup>2" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1755 | by (intro power_mono_ennreal) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1756 | also have "\<dots> = F * G" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1757 | by (simp add: algebra_simps F'_sqr G'_sqr) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1758 | finally show ?thesis | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1759 | by (simp add: F_def G_def) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1760 | qed | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1761 | qed auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1762 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1763 | |
| 69457 
bea49e443909
tagged more of HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
69313diff
changeset | 1764 | (* TODO: rename? *) | 
| 61808 | 1765 | subsection \<open>Integral under concrete measures\<close> | 
| 56994 | 1766 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1767 | lemma nn_integral_mono_measure: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1768 | assumes "sets M = sets N" "M \<le> N" shows "nn_integral M f \<le> nn_integral N f" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1769 | unfolding nn_integral_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1770 | proof (intro SUP_subset_mono) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1771 | note \<open>sets M = sets N\<close>[simp] \<open>sets M = sets N\<close>[THEN sets_eq_imp_space_eq, simp] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1772 |   show "{g. simple_function M g \<and> g \<le> f} \<subseteq> {g. simple_function N g \<and> g \<le> f}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1773 | by (simp add: simple_function_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1774 | show "integral\<^sup>S M x \<le> integral\<^sup>S N x" for x | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1775 | using le_measureD3[OF \<open>M \<le> N\<close>] | 
| 64267 | 1776 | by (auto simp add: simple_integral_def intro!: sum_mono mult_mono) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1777 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1778 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1779 | lemma nn_integral_empty: | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1780 |   assumes "space M = {}"
 | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1781 | shows "nn_integral M f = 0" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1782 | proof - | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1783 | have "(\<integral>\<^sup>+ x. f x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1784 | by(rule nn_integral_cong)(simp add: assms) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1785 | thus ?thesis by simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1786 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1787 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1788 | lemma nn_integral_bot[simp]: "nn_integral bot f = 0" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1789 | by (simp add: nn_integral_empty) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63167diff
changeset | 1790 | |
| 70136 | 1791 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Distributions\<close> | 
| 47694 | 1792 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1793 | lemma nn_integral_distr: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1794 | assumes T: "T \<in> measurable M M'" and f: "f \<in> borel_measurable (distr M M' T)" | 
| 56996 | 1795 | shows "integral\<^sup>N (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1796 | using f | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1797 | proof induct | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1798 | case (cong f g) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1799 | with T show ?case | 
| 56996 | 1800 | apply (subst nn_integral_cong[of _ f g]) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1801 | apply simp | 
| 56996 | 1802 | apply (subst nn_integral_cong[of _ "\<lambda>x. f (T x)" "\<lambda>x. g (T x)"]) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1803 | apply (simp add: measurable_def Pi_iff) | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1804 | apply simp | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1805 | done | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1806 | next | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1807 | case (set A) | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1808 | then have eq: "\<And>x. x \<in> space M \<Longrightarrow> indicator A (T x) = indicator (T -` A \<inter> space M) x" | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1809 | by (auto simp: indicator_def) | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1810 | from set T show ?case | 
| 56996 | 1811 | by (subst nn_integral_cong[OF eq]) | 
| 1812 | (auto simp add: emeasure_distr intro!: nn_integral_indicator[symmetric] measurable_sets) | |
| 1813 | qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add | |
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 1814 | nn_integral_monotone_convergence_SUP le_fun_def incseq_def image_comp) | 
| 47694 | 1815 | |
| 70136 | 1816 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Counting space\<close> | 
| 47694 | 1817 | |
| 1818 | lemma simple_function_count_space[simp]: | |
| 1819 | "simple_function (count_space A) f \<longleftrightarrow> finite (f ` A)" | |
| 1820 | unfolding simple_function_def by simp | |
| 1821 | ||
| 56996 | 1822 | lemma nn_integral_count_space: | 
| 47694 | 1823 |   assumes A: "finite {a\<in>A. 0 < f a}"
 | 
| 56996 | 1824 | shows "integral\<^sup>N (count_space A) f = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)" | 
| 35582 | 1825 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1826 | have *: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>count_space A) = | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1827 |     (\<integral>\<^sup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
 | 
| 56996 | 1828 | by (auto intro!: nn_integral_cong | 
| 73536 | 1829 | simp add: indicator_def of_bool_def if_distrib sum.If_cases[OF A] max_def le_less) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1830 |   also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^sup>+ x. f a * indicator {a} x \<partial>count_space A)"
 | 
| 64267 | 1831 | by (subst nn_integral_sum) (simp_all add: AE_count_space less_imp_le) | 
| 47694 | 1832 | also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)" | 
| 64267 | 1833 | by (auto intro!: sum.cong simp: one_ennreal_def[symmetric] max_def) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1834 | finally show ?thesis by (simp add: max.absorb2) | 
| 47694 | 1835 | qed | 
| 1836 | ||
| 56996 | 1837 | lemma nn_integral_count_space_finite: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1838 | "finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)" | 
| 64267 | 1839 | by (auto intro!: sum.mono_neutral_left simp: nn_integral_count_space less_le) | 
| 47694 | 1840 | |
| 59000 | 1841 | lemma nn_integral_count_space': | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1842 | assumes "finite A" "\<And>x. x \<in> B \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0" "A \<subseteq> B" | 
| 59000 | 1843 | shows "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>x\<in>A. f x)" | 
| 1844 | proof - | |
| 1845 | have "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>a | a \<in> B \<and> 0 < f a. f a)" | |
| 1846 | using assms(2,3) | |
| 61808 | 1847 | by (intro nn_integral_count_space finite_subset[OF _ \<open>finite A\<close>]) (auto simp: less_le) | 
| 59000 | 1848 | also have "\<dots> = (\<Sum>a\<in>A. f a)" | 
| 64267 | 1849 | using assms by (intro sum.mono_neutral_cong_left) (auto simp: less_le) | 
| 59000 | 1850 | finally show ?thesis . | 
| 1851 | qed | |
| 1852 | ||
| 59011 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1853 | lemma nn_integral_bij_count_space: | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1854 | assumes g: "bij_betw g A B" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1855 | shows "(\<integral>\<^sup>+x. f (g x) \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1856 | using g[THEN bij_betw_imp_funcset] | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1857 | by (subst distr_bij_count_space[OF g, symmetric]) | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1858 | (auto intro!: nn_integral_distr[symmetric]) | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1859 | |
| 59000 | 1860 | lemma nn_integral_indicator_finite: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1861 | fixes f :: "'a \<Rightarrow> ennreal" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1862 |   assumes f: "finite A" and [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
 | 
| 59000 | 1863 |   shows "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<Sum>x\<in>A. f x * emeasure M {x})"
 | 
| 1864 | proof - | |
| 1865 |   from f have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
 | |
| 64267 | 1866 | by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="\<lambda>a. x * a" for x] sum.If_cases) | 
| 59000 | 1867 |   also have "\<dots> = (\<Sum>a\<in>A. f a * emeasure M {a})"
 | 
| 64267 | 1868 | by (subst nn_integral_sum) auto | 
| 59000 | 1869 | finally show ?thesis . | 
| 1870 | qed | |
| 1871 | ||
| 57025 | 1872 | lemma nn_integral_count_space_nat: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1873 | fixes f :: "nat \<Rightarrow> ennreal" | 
| 57025 | 1874 | shows "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) = (\<Sum>i. f i)" | 
| 1875 | proof - | |
| 1876 | have "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) = | |
| 1877 |     (\<integral>\<^sup>+i. (\<Sum>j. f j * indicator {j} i) \<partial>count_space UNIV)"
 | |
| 1878 | proof (intro nn_integral_cong) | |
| 1879 | fix i | |
| 1880 |     have "f i = (\<Sum>j\<in>{i}. f j * indicator {j} i)"
 | |
| 1881 | by simp | |
| 1882 |     also have "\<dots> = (\<Sum>j. f j * indicator {j} i)"
 | |
| 1883 | by (rule suminf_finite[symmetric]) auto | |
| 1884 |     finally show "f i = (\<Sum>j. f j * indicator {j} i)" .
 | |
| 1885 | qed | |
| 1886 |   also have "\<dots> = (\<Sum>j. (\<integral>\<^sup>+i. f j * indicator {j} i \<partial>count_space UNIV))"
 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1887 | by (rule nn_integral_suminf) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1888 | finally show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1889 | by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1890 | qed | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1891 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1892 | lemma nn_integral_enat_function: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1893 | assumes f: "f \<in> measurable M (count_space UNIV)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1894 |   shows "(\<integral>\<^sup>+ x. ennreal_of_enat (f x) \<partial>M) = (\<Sum>t. emeasure M {x \<in> space M. t < f x})"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1895 | proof - | 
| 63040 | 1896 |   define F where "F i = {x\<in>space M. i < f x}" for i :: nat
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1897 | with assms have [measurable]: "\<And>i. F i \<in> sets M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1898 | by auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1899 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1900 |   { fix x assume "x \<in> space M"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1901 | have "(\<lambda>i::nat. if i < f x then 1 else 0) sums ennreal_of_enat (f x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1902 | using sums_If_finite[of "\<lambda>r. r < f x" "\<lambda>_. 1 :: ennreal"] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1903 | by (cases "f x") (simp_all add: sums_def of_nat_tendsto_top_ennreal) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1904 | also have "(\<lambda>i. (if i < f x then 1 else 0)) = (\<lambda>i. indicator (F i) x)" | 
| 63167 | 1905 | using \<open>x \<in> space M\<close> by (simp add: one_ennreal_def F_def fun_eq_iff) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1906 | finally have "ennreal_of_enat (f x) = (\<Sum>i. indicator (F i) x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1907 | by (simp add: sums_iff) } | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1908 | then have "(\<integral>\<^sup>+x. ennreal_of_enat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1909 | by (simp cong: nn_integral_cong) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1910 | also have "\<dots> = (\<Sum>i. emeasure M (F i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1911 | by (simp add: nn_integral_suminf) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1912 | finally show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1913 | by (simp add: F_def) | 
| 57025 | 1914 | qed | 
| 1915 | ||
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1916 | lemma nn_integral_count_space_nn_integral: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1917 | fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ennreal" | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1918 | assumes "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1919 | shows "(\<integral>\<^sup>+x. \<integral>\<^sup>+i. f i x \<partial>count_space I \<partial>M) = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. f i x \<partial>M \<partial>count_space I)" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1920 | proof cases | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1921 | assume "finite I" then show ?thesis | 
| 64267 | 1922 | by (simp add: nn_integral_count_space_finite nn_integral_sum) | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1923 | next | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1924 | assume "infinite I" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1925 |   then have [simp]: "I \<noteq> {}"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1926 | by auto | 
| 61808 | 1927 | note * = bij_betw_from_nat_into[OF \<open>countable I\<close> \<open>infinite I\<close>] | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1928 | have **: "\<And>f. (\<And>i. 0 \<le> f i) \<Longrightarrow> (\<integral>\<^sup>+i. f i \<partial>count_space I) = (\<Sum>n. f (from_nat_into I n))" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1929 | by (simp add: nn_integral_bij_count_space[symmetric, OF *] nn_integral_count_space_nat) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1930 | show ?thesis | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1931 | by (simp add: ** nn_integral_suminf from_nat_into) | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1932 | qed | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1933 | |
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63918diff
changeset | 1934 | lemma of_bool_Bex_eq_nn_integral: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63918diff
changeset | 1935 | assumes unique: "\<And>x y. x \<in> X \<Longrightarrow> y \<in> X \<Longrightarrow> P x \<Longrightarrow> P y \<Longrightarrow> x = y" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63918diff
changeset | 1936 | shows "of_bool (\<exists>y\<in>X. P y) = (\<integral>\<^sup>+y. of_bool (P y) \<partial>count_space X)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63918diff
changeset | 1937 | proof cases | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63918diff
changeset | 1938 | assume "\<exists>y\<in>X. P y" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63918diff
changeset | 1939 | then obtain y where "P y" "y \<in> X" by auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63918diff
changeset | 1940 | then show ?thesis | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63918diff
changeset | 1941 |     by (subst nn_integral_count_space'[where A="{y}"]) (auto dest: unique)
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63918diff
changeset | 1942 | qed (auto cong: nn_integral_cong_simp) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63918diff
changeset | 1943 | |
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1944 | lemma emeasure_UN_countable: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1945 | assumes sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets M" and I[simp]: "countable I" | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1946 | assumes disj: "disjoint_family_on X I" | 
| 69313 | 1947 | shows "emeasure M (\<Union>(X ` I)) = (\<integral>\<^sup>+i. emeasure M (X i) \<partial>count_space I)" | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1948 | proof - | 
| 69313 | 1949 | have eq: "\<And>x. indicator (\<Union>(X ` I)) x = \<integral>\<^sup>+ i. indicator (X i) x \<partial>count_space I" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1950 | proof cases | 
| 69313 | 1951 | fix x assume x: "x \<in> \<Union>(X ` I)" | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1952 | then obtain j where j: "x \<in> X j" "j \<in> I" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1953 | by auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1954 |     with disj have "\<And>i. i \<in> I \<Longrightarrow> indicator (X i) x = (indicator {j} i::ennreal)"
 | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1955 | by (auto simp: disjoint_family_on_def split: split_indicator) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1956 | with x j show "?thesis x" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1957 | by (simp cong: nn_integral_cong_simp) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1958 | qed (auto simp: nn_integral_0_iff_AE) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1959 | |
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1960 | note sets.countable_UN'[unfolded subset_eq, measurable] | 
| 69313 | 1961 | have "emeasure M (\<Union>(X ` I)) = (\<integral>\<^sup>+x. indicator (\<Union>(X ` I)) x \<partial>M)" | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1962 | by simp | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1963 | also have "\<dots> = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. indicator (X i) x \<partial>M \<partial>count_space I)" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1964 | by (simp add: eq nn_integral_count_space_nn_integral) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1965 | finally show ?thesis | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1966 | by (simp cong: nn_integral_cong_simp) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1967 | qed | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1968 | |
| 57025 | 1969 | lemma emeasure_countable_singleton: | 
| 1970 |   assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" and X: "countable X"
 | |
| 1971 |   shows "emeasure M X = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
 | |
| 1972 | proof - | |
| 1973 |   have "emeasure M (\<Union>i\<in>X. {i}) = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
 | |
| 1974 | using assms by (intro emeasure_UN_countable) (auto simp: disjoint_family_on_def) | |
| 1975 |   also have "(\<Union>i\<in>X. {i}) = X" by auto
 | |
| 1976 | finally show ?thesis . | |
| 1977 | qed | |
| 1978 | ||
| 1979 | lemma measure_eqI_countable: | |
| 1980 | assumes [simp]: "sets M = Pow A" "sets N = Pow A" and A: "countable A" | |
| 1981 |   assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
 | |
| 1982 | shows "M = N" | |
| 1983 | proof (rule measure_eqI) | |
| 1984 | fix X assume "X \<in> sets M" | |
| 1985 | then have X: "X \<subseteq> A" by auto | |
| 63540 | 1986 | moreover from A X have "countable X" by (auto dest: countable_subset) | 
| 57025 | 1987 | ultimately have | 
| 1988 |     "emeasure M X = (\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X)"
 | |
| 1989 |     "emeasure N X = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
 | |
| 1990 | by (auto intro!: emeasure_countable_singleton) | |
| 1991 |   moreover have "(\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X) = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
 | |
| 1992 | using X by (intro nn_integral_cong eq) auto | |
| 1993 | ultimately show "emeasure M X = emeasure N X" | |
| 1994 | by simp | |
| 1995 | qed simp | |
| 1996 | ||
| 59000 | 1997 | lemma measure_eqI_countable_AE: | 
| 1998 | assumes [simp]: "sets M = UNIV" "sets N = UNIV" | |
| 1999 | assumes ae: "AE x in M. x \<in> \<Omega>" "AE x in N. x \<in> \<Omega>" and [simp]: "countable \<Omega>" | |
| 2000 |   assumes eq: "\<And>x. x \<in> \<Omega> \<Longrightarrow> emeasure M {x} = emeasure N {x}"
 | |
| 2001 | shows "M = N" | |
| 2002 | proof (rule measure_eqI) | |
| 2003 | fix A | |
| 2004 |   have "emeasure N A = emeasure N {x\<in>\<Omega>. x \<in> A}"
 | |
| 2005 | using ae by (intro emeasure_eq_AE) auto | |
| 2006 |   also have "\<dots> = (\<integral>\<^sup>+x. emeasure N {x} \<partial>count_space {x\<in>\<Omega>. x \<in> A})"
 | |
| 2007 | by (intro emeasure_countable_singleton) auto | |
| 2008 |   also have "\<dots> = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space {x\<in>\<Omega>. x \<in> A})"
 | |
| 2009 | by (intro nn_integral_cong eq[symmetric]) auto | |
| 2010 |   also have "\<dots> = emeasure M {x\<in>\<Omega>. x \<in> A}"
 | |
| 2011 | by (intro emeasure_countable_singleton[symmetric]) auto | |
| 2012 | also have "\<dots> = emeasure M A" | |
| 2013 | using ae by (intro emeasure_eq_AE) auto | |
| 2014 | finally show "emeasure M A = emeasure N A" .. | |
| 2015 | qed simp | |
| 2016 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2017 | lemma nn_integral_monotone_convergence_SUP_nat: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2018 | fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal" | 
| 67399 | 2019 | assumes chain: "Complete_Partial_Order.chain (\<le>) (f ` Y)" | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2020 |   and nonempty: "Y \<noteq> {}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2021 | shows "(\<integral>\<^sup>+ x. (SUP i\<in>Y. f i x) \<partial>count_space UNIV) = (SUP i\<in>Y. (\<integral>\<^sup>+ x. f i x \<partial>count_space UNIV))" | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2022 | (is "?lhs = ?rhs" is "integral\<^sup>N ?M _ = _") | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2023 | proof (rule order_class.order.antisym) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2024 | show "?rhs \<le> ?lhs" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2025 | by (auto intro!: SUP_least SUP_upper nn_integral_mono) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2026 | next | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2027 | have "\<exists>g. incseq g \<and> range g \<subseteq> (\<lambda>i. f i x) ` Y \<and> (SUP i\<in>Y. f i x) = (SUP i. g i)" for x | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2028 | by (rule ennreal_Sup_countable_SUP) (simp add: nonempty) | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2029 | then obtain g where incseq: "\<And>x. incseq (g x)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2030 | and range: "\<And>x. range (g x) \<subseteq> (\<lambda>i. f i x) ` Y" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2031 | and sup: "\<And>x. (SUP i\<in>Y. f i x) = (SUP i. g x i)" by moura | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2032 | from incseq have incseq': "incseq (\<lambda>i x. g x i)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2033 | by(blast intro: incseq_SucI le_funI dest: incseq_SucD) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2034 | |
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2035 | have "?lhs = \<integral>\<^sup>+ x. (SUP i. g x i) \<partial>?M" by(simp add: sup) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2036 | also have "\<dots> = (SUP i. \<integral>\<^sup>+ x. g x i \<partial>?M)" using incseq' | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2037 | by(rule nn_integral_monotone_convergence_SUP) simp | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2038 | also have "\<dots> \<le> (SUP i\<in>Y. \<integral>\<^sup>+ x. f i x \<partial>?M)" | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2039 | proof(rule SUP_least) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2040 | fix n | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2041 | have "\<And>x. \<exists>i. g x n = f i x \<and> i \<in> Y" using range by blast | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2042 | then obtain I where I: "\<And>x. g x n = f (I x) x" "\<And>x. I x \<in> Y" by moura | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2043 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2044 | have "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) = (\<Sum>x. g x n)" | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2045 | by(rule nn_integral_count_space_nat) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2046 | also have "\<dots> = (SUP m. \<Sum>x<m. g x n)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2047 | by(rule suminf_eq_SUP) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2048 | also have "\<dots> \<le> (SUP i\<in>Y. \<integral>\<^sup>+ x. f i x \<partial>?M)" | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2049 | proof(rule SUP_mono) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2050 | fix m | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2051 | show "\<exists>m'\<in>Y. (\<Sum>x<m. g x n) \<le> (\<integral>\<^sup>+ x. f m' x \<partial>?M)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2052 | proof(cases "m > 0") | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2053 | case False | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2054 | thus ?thesis using nonempty by auto | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2055 | next | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2056 | case True | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2057 |         let ?Y = "I ` {..<m}"
 | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2058 | have "f ` ?Y \<subseteq> f ` Y" using I by auto | 
| 67399 | 2059 | with chain have chain': "Complete_Partial_Order.chain (\<le>) (f ` ?Y)" by(rule chain_subset) | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2060 | hence "Sup (f ` ?Y) \<in> f ` ?Y" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2061 | by(rule ccpo_class.in_chain_finite)(auto simp add: True lessThan_empty_iff) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2062 | then obtain m' where "m' < m" and m': "(SUP i\<in>?Y. f i) = f (I m')" by auto | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2063 | have "I m' \<in> Y" using I by blast | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2064 | have "(\<Sum>x<m. g x n) \<le> (\<Sum>x<m. f (I m') x)" | 
| 64267 | 2065 | proof(rule sum_mono) | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2066 | fix x | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2067 |           assume "x \<in> {..<m}"
 | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2068 | hence "x < m" by simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2069 | have "g x n = f (I x) x" by(simp add: I) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2070 | also have "\<dots> \<le> (SUP i\<in>?Y. f i) x" unfolding Sup_fun_def image_image | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62083diff
changeset | 2071 |             using \<open>x \<in> {..<m}\<close> by (rule Sup_upper [OF imageI])
 | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2072 | also have "\<dots> = f (I m') x" unfolding m' by simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2073 | finally show "g x n \<le> f (I m') x" . | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2074 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2075 | also have "\<dots> \<le> (SUP m. (\<Sum>x<m. f (I m') x))" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2076 | by(rule SUP_upper) simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2077 | also have "\<dots> = (\<Sum>x. f (I m') x)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2078 | by(rule suminf_eq_SUP[symmetric]) | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2079 | also have "\<dots> = (\<integral>\<^sup>+ x. f (I m') x \<partial>?M)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2080 | by(rule nn_integral_count_space_nat[symmetric]) | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2081 | finally show ?thesis using \<open>I m' \<in> Y\<close> by blast | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2082 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2083 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2084 | finally show "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) \<le> \<dots>" . | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2085 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2086 | finally show "?lhs \<le> ?rhs" . | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2087 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2088 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2089 | lemma power_series_tendsto_at_left: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2090 | assumes nonneg: "\<And>i. 0 \<le> f i" and summable: "\<And>z. 0 \<le> z \<Longrightarrow> z < 1 \<Longrightarrow> summable (\<lambda>n. f n * z^n)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2091 | shows "((\<lambda>z. ennreal (\<Sum>n. f n * z^n)) \<longlongrightarrow> (\<Sum>n. ennreal (f n))) (at_left (1::real))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2092 | proof (intro tendsto_at_left_sequentially) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2093 | show "0 < (1::real)" by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2094 | fix S :: "nat \<Rightarrow> real" assume S: "\<And>n. S n < 1" "\<And>n. 0 < S n" "S \<longlonglongrightarrow> 1" "incseq S" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2095 | then have S_nonneg: "\<And>i. 0 \<le> S i" by (auto intro: less_imp_le) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2096 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2097 | have "(\<lambda>i. (\<integral>\<^sup>+n. f n * S i^n \<partial>count_space UNIV)) \<longlonglongrightarrow> (\<integral>\<^sup>+n. ennreal (f n) \<partial>count_space UNIV)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2098 | proof (rule nn_integral_LIMSEQ) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2099 | show "incseq (\<lambda>i n. ennreal (f n * S i^n))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2100 | using S by (auto intro!: mult_mono power_mono nonneg ennreal_leI | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2101 | simp: incseq_def le_fun_def less_imp_le) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2102 | fix n have "(\<lambda>i. ennreal (f n * S i^n)) \<longlonglongrightarrow> ennreal (f n * 1^n)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2103 | by (intro tendsto_intros tendsto_ennrealI S) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2104 | then show "(\<lambda>i. ennreal (f n * S i^n)) \<longlonglongrightarrow> ennreal (f n)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2105 | by simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2106 | qed (auto simp: S_nonneg intro!: mult_nonneg_nonneg nonneg) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2107 | also have "(\<lambda>i. (\<integral>\<^sup>+n. f n * S i^n \<partial>count_space UNIV)) = (\<lambda>i. \<Sum>n. f n * S i^n)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2108 | by (subst nn_integral_count_space_nat) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2109 | (intro ext suminf_ennreal2 mult_nonneg_nonneg nonneg S_nonneg | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2110 | zero_le_power summable S)+ | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2111 | also have "(\<integral>\<^sup>+n. ennreal (f n) \<partial>count_space UNIV) = (\<Sum>n. ennreal (f n))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2112 | by (simp add: nn_integral_count_space_nat nonneg) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2113 | finally show "(\<lambda>n. ennreal (\<Sum>na. f na * S n ^ na)) \<longlonglongrightarrow> (\<Sum>n. ennreal (f n))" . | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2114 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2115 | |
| 61808 | 2116 | subsubsection \<open>Measures with Restricted Space\<close> | 
| 54417 | 2117 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2118 | lemma simple_function_restrict_space_ennreal: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2119 | fixes f :: "'a \<Rightarrow> ennreal" | 
| 57137 | 2120 | assumes "\<Omega> \<inter> space M \<in> sets M" | 
| 2121 | shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 2122 | proof - | |
| 2123 |   { assume "finite (f ` space (restrict_space M \<Omega>))"
 | |
| 2124 |     then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
 | |
| 2125 | then have "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)" | |
| 2126 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 2127 | moreover | |
| 2128 |   { assume "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)"
 | |
| 2129 | then have "finite (f ` space (restrict_space M \<Omega>))" | |
| 2130 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 2131 | ultimately show ?thesis | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2132 | unfolding | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2133 | simple_function_iff_borel_measurable borel_measurable_restrict_space_iff_ennreal[OF assms] | 
| 57137 | 2134 | by auto | 
| 2135 | qed | |
| 2136 | ||
| 2137 | lemma simple_function_restrict_space: | |
| 2138 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | |
| 2139 | assumes "\<Omega> \<inter> space M \<in> sets M" | |
| 2140 | shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" | |
| 2141 | proof - | |
| 2142 |   { assume "finite (f ` space (restrict_space M \<Omega>))"
 | |
| 2143 |     then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
 | |
| 2144 | then have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)" | |
| 2145 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 2146 | moreover | |
| 2147 |   { assume "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)"
 | |
| 2148 | then have "finite (f ` space (restrict_space M \<Omega>))" | |
| 2149 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 2150 | ultimately show ?thesis | |
| 2151 | unfolding simple_function_iff_borel_measurable | |
| 2152 | borel_measurable_restrict_space_iff[OF assms] | |
| 2153 | by auto | |
| 2154 | qed | |
| 2155 | ||
| 2156 | lemma simple_integral_restrict_space: | |
| 2157 | assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" "simple_function (restrict_space M \<Omega>) f" | |
| 2158 | shows "simple_integral (restrict_space M \<Omega>) f = simple_integral M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2159 | using simple_function_restrict_space_ennreal[THEN iffD1, OF \<Omega>, THEN simple_functionD(1)] | 
| 57137 | 2160 | by (auto simp add: space_restrict_space emeasure_restrict_space[OF \<Omega>(1)] le_infI2 simple_integral_def | 
| 2161 | split: split_indicator split_indicator_asm | |
| 64267 | 2162 | intro!: sum.mono_neutral_cong_left ennreal_mult_left_cong arg_cong2[where f=emeasure]) | 
| 57137 | 2163 | |
| 56996 | 2164 | lemma nn_integral_restrict_space: | 
| 57137 | 2165 | assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M" | 
| 2166 | shows "nn_integral (restrict_space M \<Omega>) f = nn_integral M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 2167 | proof - | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2168 |   let ?R = "restrict_space M \<Omega>" and ?X = "\<lambda>M f. {s. simple_function M s \<and> s \<le> f \<and> (\<forall>x. s x < top)}"
 | 
| 57137 | 2169 | have "integral\<^sup>S ?R ` ?X ?R f = integral\<^sup>S M ` ?X M (\<lambda>x. f x * indicator \<Omega> x)" | 
| 2170 | proof (safe intro!: image_eqI) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2171 | fix s assume s: "simple_function ?R s" "s \<le> f" "\<forall>x. s x < top" | 
| 57137 | 2172 | from s show "integral\<^sup>S (restrict_space M \<Omega>) s = integral\<^sup>S M (\<lambda>x. s x * indicator \<Omega> x)" | 
| 2173 | by (intro simple_integral_restrict_space) auto | |
| 2174 | from s show "simple_function M (\<lambda>x. s x * indicator \<Omega> x)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2175 | by (simp add: simple_function_restrict_space_ennreal) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2176 | from s show "(\<lambda>x. s x * indicator \<Omega> x) \<le> (\<lambda>x. f x * indicator \<Omega> x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2177 | "\<And>x. s x * indicator \<Omega> x < top" | 
| 57137 | 2178 | by (auto split: split_indicator simp: le_fun_def image_subset_iff) | 
| 2179 | next | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2180 | fix s assume s: "simple_function M s" "s \<le> (\<lambda>x. f x * indicator \<Omega> x)" "\<forall>x. s x < top" | 
| 57137 | 2181 | then have "simple_function M (\<lambda>x. s x * indicator (\<Omega> \<inter> space M) x)" (is ?s') | 
| 2182 | by (intro simple_function_mult simple_function_indicator) auto | |
| 2183 | also have "?s' \<longleftrightarrow> simple_function M (\<lambda>x. s x * indicator \<Omega> x)" | |
| 2184 | by (rule simple_function_cong) (auto split: split_indicator) | |
| 2185 | finally show sf: "simple_function (restrict_space M \<Omega>) s" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2186 | by (simp add: simple_function_restrict_space_ennreal) | 
| 57137 | 2187 | |
| 2188 | from s have s_eq: "s = (\<lambda>x. s x * indicator \<Omega> x)" | |
| 2189 | by (auto simp add: fun_eq_iff le_fun_def image_subset_iff | |
| 2190 | split: split_indicator split_indicator_asm | |
| 2191 | intro: antisym) | |
| 2192 | ||
| 2193 | show "integral\<^sup>S M s = integral\<^sup>S (restrict_space M \<Omega>) s" | |
| 2194 | by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \<Omega> sf]) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2195 | show "\<And>x. s x < top" | 
| 57137 | 2196 | using s by (auto simp: image_subset_iff) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2197 | from s show "s \<le> f" | 
| 57137 | 2198 | by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm) | 
| 2199 | qed | |
| 2200 | then show ?thesis | |
| 69546 
27dae626822b
prefer naming convention from datatype package for strong congruence rules
 haftmann parents: 
69457diff
changeset | 2201 | unfolding nn_integral_def_finite by (simp cong del: SUP_cong_simp) | 
| 54417 | 2202 | qed | 
| 2203 | ||
| 59000 | 2204 | lemma nn_integral_count_space_indicator: | 
| 59779 | 2205 | assumes "NO_MATCH (UNIV::'a set) (X::'a set)" | 
| 59000 | 2206 | shows "(\<integral>\<^sup>+x. f x \<partial>count_space X) = (\<integral>\<^sup>+x. f x * indicator X x \<partial>count_space UNIV)" | 
| 2207 | by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space) | |
| 2208 | ||
| 59425 | 2209 | lemma nn_integral_count_space_eq: | 
| 2210 | "(\<And>x. x \<in> A - B \<Longrightarrow> f x = 0) \<Longrightarrow> (\<And>x. x \<in> B - A \<Longrightarrow> f x = 0) \<Longrightarrow> | |
| 2211 | (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)" | |
| 2212 | by (auto simp: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator) | |
| 2213 | ||
| 59023 | 2214 | lemma nn_integral_ge_point: | 
| 2215 | assumes "x \<in> A" | |
| 2216 | shows "p x \<le> \<integral>\<^sup>+ x. p x \<partial>count_space A" | |
| 2217 | proof - | |
| 2218 |   from assms have "p x \<le> \<integral>\<^sup>+ x. p x \<partial>count_space {x}"
 | |
| 2219 | by(auto simp add: nn_integral_count_space_finite max_def) | |
| 2220 |   also have "\<dots> = \<integral>\<^sup>+ x'. p x' * indicator {x} x' \<partial>count_space A"
 | |
| 2221 | using assms by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2222 | also have "\<dots> \<le> \<integral>\<^sup>+ x. p x \<partial>count_space A" | 
| 59023 | 2223 | by(rule nn_integral_mono)(simp add: indicator_def) | 
| 2224 | finally show ?thesis . | |
| 2225 | qed | |
| 2226 | ||
| 61808 | 2227 | subsubsection \<open>Measure spaces with an associated density\<close> | 
| 47694 | 2228 | |
| 70136 | 2229 | definition\<^marker>\<open>tag important\<close> density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2230 | "density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M)" | 
| 35582 | 2231 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2232 | lemma | 
| 59048 | 2233 | shows sets_density[simp, measurable_cong]: "sets (density M f) = sets M" | 
| 47694 | 2234 | and space_density[simp]: "space (density M f) = space M" | 
| 2235 | by (auto simp: density_def) | |
| 2236 | ||
| 50003 | 2237 | (* FIXME: add conversion to simplify space, sets and measurable *) | 
| 2238 | lemma space_density_imp[measurable_dest]: | |
| 2239 | "\<And>x M f. x \<in> space (density M f) \<Longrightarrow> x \<in> space M" by auto | |
| 2240 | ||
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2241 | lemma | 
| 47694 | 2242 | shows measurable_density_eq1[simp]: "g \<in> measurable (density Mg f) Mg' \<longleftrightarrow> g \<in> measurable Mg Mg'" | 
| 2243 | and measurable_density_eq2[simp]: "h \<in> measurable Mh (density Mh' f) \<longleftrightarrow> h \<in> measurable Mh Mh'" | |
| 2244 | and simple_function_density_eq[simp]: "simple_function (density Mu f) u \<longleftrightarrow> simple_function Mu u" | |
| 2245 | unfolding measurable_def simple_function_def by simp_all | |
| 2246 | ||
| 2247 | lemma density_cong: "f \<in> borel_measurable M \<Longrightarrow> f' \<in> borel_measurable M \<Longrightarrow> | |
| 2248 | (AE x in M. f x = f' x) \<Longrightarrow> density M f = density M f'" | |
| 56996 | 2249 | unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed) | 
| 47694 | 2250 | |
| 2251 | lemma emeasure_density: | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 2252 | assumes f[measurable]: "f \<in> borel_measurable M" and A[measurable]: "A \<in> sets M" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2253 | shows "emeasure (density M f) A = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M)" | 
| 47694 | 2254 | (is "_ = ?\<mu> A") | 
| 2255 | unfolding density_def | |
| 2256 | proof (rule emeasure_measure_of_sigma) | |
| 2257 | show "sigma_algebra (space M) (sets M)" .. | |
| 2258 | show "positive (sets M) ?\<mu>" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2259 | using f by (auto simp: positive_def) | 
| 47694 | 2260 | show "countably_additive (sets M) ?\<mu>" | 
| 2261 | proof (intro countably_additiveI) | |
| 2262 | fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M" | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 2263 | then have "\<And>i. A i \<in> sets M" by auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2264 | then have *: "\<And>i. (\<lambda>x. f x * indicator (A i) x) \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2265 | by auto | 
| 47694 | 2266 | assume disj: "disjoint_family A" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2267 | then have "(\<Sum>n. ?\<mu> (A n)) = (\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (A n) x) \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2268 | using f * by (subst nn_integral_suminf) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2269 | also have "(\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (A n) x) \<partial>M) = (\<integral>\<^sup>+ x. f x * (\<Sum>n. indicator (A n) x) \<partial>M)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2270 | using f by (auto intro!: ennreal_suminf_cmult nn_integral_cong_AE) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2271 | also have "\<dots> = (\<integral>\<^sup>+ x. f x * indicator (\<Union>n. A n) x \<partial>M)" | 
| 47694 | 2272 | unfolding suminf_indicator[OF disj] .. | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2273 | finally show "(\<Sum>i. \<integral>\<^sup>+ x. f x * indicator (A i) x \<partial>M) = \<integral>\<^sup>+ x. f x * indicator (\<Union>i. A i) x \<partial>M" . | 
| 47694 | 2274 | qed | 
| 2275 | qed fact | |
| 38656 | 2276 | |
| 47694 | 2277 | lemma null_sets_density_iff: | 
| 2278 | assumes f: "f \<in> borel_measurable M" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2279 | shows "A \<in> null_sets (density M f) \<longleftrightarrow> A \<in> sets M \<and> (AE x in M. x \<in> A \<longrightarrow> f x = 0)" | 
| 47694 | 2280 | proof - | 
| 2281 |   { assume "A \<in> sets M"
 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2282 |     have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> emeasure M {x \<in> space M. f x * indicator A x \<noteq> 0} = 0"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2283 | using f \<open>A \<in> sets M\<close> by (intro nn_integral_0_iff) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2284 | also have "\<dots> \<longleftrightarrow> (AE x in M. f x * indicator A x = 0)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2285 | using f \<open>A \<in> sets M\<close> by (intro AE_iff_measurable[OF _ refl, symmetric]) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2286 | also have "(AE x in M. f x * indicator A x = 0) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" | 
| 62390 | 2287 | by (auto simp add: indicator_def max_def split: if_split_asm) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2288 | finally have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" . } | 
| 47694 | 2289 | with f show ?thesis | 
| 2290 | by (simp add: null_sets_def emeasure_density cong: conj_cong) | |
| 2291 | qed | |
| 2292 | ||
| 2293 | lemma AE_density: | |
| 2294 | assumes f: "f \<in> borel_measurable M" | |
| 2295 | shows "(AE x in density M f. P x) \<longleftrightarrow> (AE x in M. 0 < f x \<longrightarrow> P x)" | |
| 2296 | proof | |
| 2297 | assume "AE x in density M f. P x" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2298 |   with f obtain N where "{x \<in> space M. \<not> P x} \<subseteq> N" "N \<in> sets M" and ae: "AE x in M. x \<in> N \<longrightarrow> f x = 0"
 | 
| 47694 | 2299 | by (auto simp: eventually_ae_filter null_sets_density_iff) | 
| 2300 | then have "AE x in M. x \<notin> N \<longrightarrow> P x" by auto | |
| 2301 | with ae show "AE x in M. 0 < f x \<longrightarrow> P x" | |
| 2302 | by (rule eventually_elim2) auto | |
| 2303 | next | |
| 2304 | fix N assume ae: "AE x in M. 0 < f x \<longrightarrow> P x" | |
| 2305 |   then obtain N where "{x \<in> space M. \<not> (0 < f x \<longrightarrow> P x)} \<subseteq> N" "N \<in> null_sets M"
 | |
| 2306 | by (auto simp: eventually_ae_filter) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2307 |   then have *: "{x \<in> space (density M f). \<not> P x} \<subseteq> N \<union> {x\<in>space M. f x = 0}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2308 |     "N \<union> {x\<in>space M. f x = 0} \<in> sets M" and ae2: "AE x in M. x \<notin> N"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2309 | using f by (auto simp: subset_eq zero_less_iff_neq_zero intro!: AE_not_in) | 
| 47694 | 2310 | show "AE x in density M f. P x" | 
| 2311 | using ae2 | |
| 2312 | unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f] | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2313 |     by (intro exI[of _ "N \<union> {x\<in>space M. f x = 0}"] conjI *) (auto elim: eventually_elim2)
 | 
| 35582 | 2314 | qed | 
| 2315 | ||
| 70136 | 2316 | lemma\<^marker>\<open>tag important\<close> nn_integral_density: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2317 | assumes f: "f \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2318 | assumes g: "g \<in> borel_measurable M" | 
| 56996 | 2319 | shows "integral\<^sup>N (density M f) g = (\<integral>\<^sup>+ x. f x * g x \<partial>M)" | 
| 70136 | 2320 | using g proof induct | 
| 49798 | 2321 | case (cong u v) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 2322 | then show ?case | 
| 56996 | 2323 | apply (subst nn_integral_cong[OF cong(3)]) | 
| 2324 | apply (simp_all cong: nn_integral_cong) | |
| 49798 | 2325 | done | 
| 2326 | next | |
| 2327 | case (set A) then show ?case | |
| 2328 | by (simp add: emeasure_density f) | |
| 2329 | next | |
| 2330 | case (mult u c) | |
| 2331 | moreover have "\<And>x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps) | |
| 2332 | ultimately show ?case | |
| 56996 | 2333 | using f by (simp add: nn_integral_cmult) | 
| 49798 | 2334 | next | 
| 2335 | case (add u v) | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 2336 | then have "\<And>x. f x * (v x + u x) = f x * v x + f x * u x" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2337 | by (simp add: distrib_left) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 2338 | with add f show ?case | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2339 | by (auto simp add: nn_integral_add intro!: nn_integral_add[symmetric]) | 
| 49798 | 2340 | next | 
| 2341 | case (seq U) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2342 | have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2343 | by eventually_elim (simp add: SUP_mult_left_ennreal seq) | 
| 49798 | 2344 | from seq f show ?case | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 2345 | apply (simp add: nn_integral_monotone_convergence_SUP image_comp) | 
| 56996 | 2346 | apply (subst nn_integral_cong_AE[OF eq]) | 
| 2347 | apply (subst nn_integral_monotone_convergence_SUP_AE) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2348 | apply (auto simp: incseq_def le_fun_def intro!: mult_left_mono) | 
| 49798 | 2349 | done | 
| 47694 | 2350 | qed | 
| 38705 | 2351 | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2352 | lemma density_distr: | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2353 | assumes [measurable]: "f \<in> borel_measurable N" "X \<in> measurable M N" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2354 | shows "density (distr M N X) f = distr (density M (\<lambda>x. f (X x))) N X" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2355 | by (intro measure_eqI) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2356 | (auto simp add: emeasure_density nn_integral_distr emeasure_distr | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2357 | split: split_indicator intro!: nn_integral_cong) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2358 | |
| 47694 | 2359 | lemma emeasure_restricted: | 
| 2360 | assumes S: "S \<in> sets M" and X: "X \<in> sets M" | |
| 2361 | shows "emeasure (density M (indicator S)) X = emeasure M (S \<inter> X)" | |
| 38705 | 2362 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2363 | have "emeasure (density M (indicator S)) X = (\<integral>\<^sup>+x. indicator S x * indicator X x \<partial>M)" | 
| 47694 | 2364 | using S X by (simp add: emeasure_density) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2365 | also have "\<dots> = (\<integral>\<^sup>+x. indicator (S \<inter> X) x \<partial>M)" | 
| 56996 | 2366 | by (auto intro!: nn_integral_cong simp: indicator_def) | 
| 47694 | 2367 | also have "\<dots> = emeasure M (S \<inter> X)" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 2368 | using S X by (simp add: sets.Int) | 
| 47694 | 2369 | finally show ?thesis . | 
| 2370 | qed | |
| 2371 | ||
| 2372 | lemma measure_restricted: | |
| 2373 | "S \<in> sets M \<Longrightarrow> X \<in> sets M \<Longrightarrow> measure (density M (indicator S)) X = measure M (S \<inter> X)" | |
| 2374 | by (simp add: emeasure_restricted measure_def) | |
| 2375 | ||
| 2376 | lemma (in finite_measure) finite_measure_restricted: | |
| 2377 | "S \<in> sets M \<Longrightarrow> finite_measure (density M (indicator S))" | |
| 61169 | 2378 | by standard (simp add: emeasure_restricted) | 
| 47694 | 2379 | |
| 2380 | lemma emeasure_density_const: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2381 | "A \<in> sets M \<Longrightarrow> emeasure (density M (\<lambda>_. c)) A = c * emeasure M A" | 
| 56996 | 2382 | by (auto simp: nn_integral_cmult_indicator emeasure_density) | 
| 47694 | 2383 | |
| 2384 | lemma measure_density_const: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2385 | "A \<in> sets M \<Longrightarrow> c \<noteq> \<infinity> \<Longrightarrow> measure (density M (\<lambda>_. c)) A = enn2real c * measure M A" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2386 | by (auto simp: emeasure_density_const measure_def enn2real_mult) | 
| 47694 | 2387 | |
| 2388 | lemma density_density_eq: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2389 | "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> | 
| 47694 | 2390 | density (density M f) g = density M (\<lambda>x. f x * g x)" | 
| 56996 | 2391 | by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps) | 
| 47694 | 2392 | |
| 2393 | lemma distr_density_distr: | |
| 2394 | assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M" | |
| 2395 | and inv: "\<forall>x\<in>space M. T' (T x) = x" | |
| 2396 | assumes f: "f \<in> borel_measurable M'" | |
| 2397 | shows "distr (density (distr M M' T) f) M T' = density M (f \<circ> T)" (is "?R = ?L") | |
| 2398 | proof (rule measure_eqI) | |
| 2399 | fix A assume A: "A \<in> sets ?R" | |
| 2400 |   { fix x assume "x \<in> space M"
 | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 2401 | with sets.sets_into_space[OF A] | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2402 | have "indicator (T' -` A \<inter> space M') (T x) = (indicator A x :: ennreal)" | 
| 47694 | 2403 | using T inv by (auto simp: indicator_def measurable_space) } | 
| 2404 | with A T T' f show "emeasure ?R A = emeasure ?L A" | |
| 2405 | by (simp add: measurable_comp emeasure_density emeasure_distr | |
| 56996 | 2406 | nn_integral_distr measurable_sets cong: nn_integral_cong) | 
| 47694 | 2407 | qed simp | 
| 2408 | ||
| 2409 | lemma density_density_divide: | |
| 2410 | fixes f g :: "'a \<Rightarrow> real" | |
| 2411 | assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" | |
| 2412 | assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" | |
| 2413 | assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0" | |
| 2414 | shows "density (density M f) (\<lambda>x. g x / f x) = density M g" | |
| 2415 | proof - | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2416 | have "density M g = density M (\<lambda>x. ennreal (f x) * ennreal (g x / f x))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2417 | using f g ac by (auto intro!: density_cong measurable_If simp: ennreal_mult[symmetric]) | 
| 47694 | 2418 | then show ?thesis | 
| 2419 | using f g by (subst density_density_eq) auto | |
| 38705 | 2420 | qed | 
| 2421 | ||
| 59425 | 2422 | lemma density_1: "density M (\<lambda>_. 1) = M" | 
| 2423 | by (intro measure_eqI) (auto simp: emeasure_density) | |
| 2424 | ||
| 2425 | lemma emeasure_density_add: | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2426 | assumes X: "X \<in> sets M" | 
| 59425 | 2427 | assumes Mf[measurable]: "f \<in> borel_measurable M" | 
| 2428 | assumes Mg[measurable]: "g \<in> borel_measurable M" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2429 | shows "emeasure (density M f) X + emeasure (density M g) X = | 
| 59425 | 2430 | emeasure (density M (\<lambda>x. f x + g x)) X" | 
| 2431 | using assms | |
| 2432 | apply (subst (1 2 3) emeasure_density, simp_all) [] | |
| 2433 | apply (subst nn_integral_add[symmetric], simp_all) [] | |
| 2434 | apply (intro nn_integral_cong, simp split: split_indicator) | |
| 2435 | done | |
| 2436 | ||
| 61808 | 2437 | subsubsection \<open>Point measure\<close> | 
| 47694 | 2438 | |
| 70136 | 2439 | definition\<^marker>\<open>tag important\<close> point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
 | 
| 47694 | 2440 | "point_measure A f = density (count_space A) f" | 
| 2441 | ||
| 2442 | lemma | |
| 2443 | shows space_point_measure: "space (point_measure A f) = A" | |
| 2444 | and sets_point_measure: "sets (point_measure A f) = Pow A" | |
| 2445 | by (auto simp: point_measure_def) | |
| 2446 | ||
| 59048 | 2447 | lemma sets_point_measure_count_space[measurable_cong]: "sets (point_measure A f) = sets (count_space A)" | 
| 2448 | by (simp add: sets_point_measure) | |
| 2449 | ||
| 47694 | 2450 | lemma measurable_point_measure_eq1[simp]: | 
| 2451 | "g \<in> measurable (point_measure A f) M \<longleftrightarrow> g \<in> A \<rightarrow> space M" | |
| 2452 | unfolding point_measure_def by simp | |
| 2453 | ||
| 2454 | lemma measurable_point_measure_eq2_finite[simp]: | |
| 2455 | "finite A \<Longrightarrow> | |
| 2456 | g \<in> measurable M (point_measure A f) \<longleftrightarrow> | |
| 2457 |     (g \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. g -` {a} \<inter> space M \<in> sets M))"
 | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 2458 | unfolding point_measure_def by (simp add: measurable_count_space_eq2) | 
| 47694 | 2459 | |
| 2460 | lemma simple_function_point_measure[simp]: | |
| 2461 | "simple_function (point_measure A f) g \<longleftrightarrow> finite (g ` A)" | |
| 2462 | by (simp add: point_measure_def) | |
| 2463 | ||
| 2464 | lemma emeasure_point_measure: | |
| 2465 |   assumes A: "finite {a\<in>X. 0 < f a}" "X \<subseteq> A"
 | |
| 2466 | shows "emeasure (point_measure A f) X = (\<Sum>a|a\<in>X \<and> 0 < f a. f a)" | |
| 35977 | 2467 | proof - | 
| 47694 | 2468 |   have "{a. (a \<in> X \<longrightarrow> a \<in> A \<and> 0 < f a) \<and> a \<in> X} = {a\<in>X. 0 < f a}"
 | 
| 61808 | 2469 | using \<open>X \<subseteq> A\<close> by auto | 
| 47694 | 2470 | with A show ?thesis | 
| 73536 | 2471 | by (simp add: emeasure_density nn_integral_count_space point_measure_def indicator_def of_bool_def) | 
| 35977 | 2472 | qed | 
| 2473 | ||
| 47694 | 2474 | lemma emeasure_point_measure_finite: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2475 | "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)" | 
| 64267 | 2476 | by (subst emeasure_point_measure) (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le) | 
| 47694 | 2477 | |
| 79583 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2478 | lemma emeasure_point_measure_finite_if: | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2479 | "finite A \<Longrightarrow> emeasure (point_measure A f) X = (if X \<subseteq> A then \<Sum>a\<in>X. f a else 0)" | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2480 | by (simp add: emeasure_point_measure_finite emeasure_notin_sets sets_point_measure) | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2481 | |
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2482 | lemma measure_point_measure_finite_if: | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2483 | assumes "finite A" and "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> 0" | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2484 | shows "measure (point_measure A f) X = (if X \<subseteq> A then \<Sum>a\<in>X. f a else 0)" | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2485 | by (simp add: Sigma_Algebra.measure_def assms emeasure_point_measure_finite_if subset_eq sum_nonneg) | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2486 | |
| 49795 | 2487 | lemma emeasure_point_measure_finite2: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2488 | "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)" | 
| 49795 | 2489 | by (subst emeasure_point_measure) | 
| 64267 | 2490 | (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le) | 
| 49795 | 2491 | |
| 47694 | 2492 | lemma null_sets_point_measure_iff: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2493 | "X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x = 0)" | 
| 47694 | 2494 | by (auto simp: AE_count_space null_sets_density_iff point_measure_def) | 
| 2495 | ||
| 2496 | lemma AE_point_measure: | |
| 2497 | "(AE x in point_measure A f. P x) \<longleftrightarrow> (\<forall>x\<in>A. 0 < f x \<longrightarrow> P x)" | |
| 2498 | unfolding point_measure_def | |
| 2499 | by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def) | |
| 2500 | ||
| 56996 | 2501 | lemma nn_integral_point_measure: | 
| 47694 | 2502 |   "finite {a\<in>A. 0 < f a \<and> 0 < g a} \<Longrightarrow>
 | 
| 56996 | 2503 | integral\<^sup>N (point_measure A f) g = (\<Sum>a|a\<in>A \<and> 0 < f a \<and> 0 < g a. f a * g a)" | 
| 47694 | 2504 | unfolding point_measure_def | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2505 | by (subst nn_integral_density) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2506 | (simp_all add: nn_integral_density nn_integral_count_space ennreal_zero_less_mult_iff) | 
| 47694 | 2507 | |
| 56996 | 2508 | lemma nn_integral_point_measure_finite: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2509 | "finite A \<Longrightarrow> integral\<^sup>N (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)" | 
| 64267 | 2510 | by (subst nn_integral_point_measure) (auto intro!: sum.mono_neutral_left simp: less_le) | 
| 47694 | 2511 | |
| 61808 | 2512 | subsubsection \<open>Uniform measure\<close> | 
| 47694 | 2513 | |
| 70136 | 2514 | definition\<^marker>\<open>tag important\<close> "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)" | 
| 47694 | 2515 | |
| 2516 | lemma | |
| 59048 | 2517 | shows sets_uniform_measure[simp, measurable_cong]: "sets (uniform_measure M A) = sets M" | 
| 47694 | 2518 | and space_uniform_measure[simp]: "space (uniform_measure M A) = space M" | 
| 2519 | by (auto simp: uniform_measure_def) | |
| 2520 | ||
| 2521 | lemma emeasure_uniform_measure[simp]: | |
| 2522 | assumes A: "A \<in> sets M" and B: "B \<in> sets M" | |
| 2523 | shows "emeasure (uniform_measure M A) B = emeasure M (A \<inter> B) / emeasure M A" | |
| 2524 | proof - | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2525 | from A B have "emeasure (uniform_measure M A) B = (\<integral>\<^sup>+x. (1 / emeasure M A) * indicator (A \<inter> B) x \<partial>M)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2526 | by (auto simp add: uniform_measure_def emeasure_density divide_ennreal_def split: split_indicator | 
| 56996 | 2527 | intro!: nn_integral_cong) | 
| 47694 | 2528 | also have "\<dots> = emeasure M (A \<inter> B) / emeasure M A" | 
| 2529 | using A B | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2530 | by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int divide_ennreal_def mult.commute) | 
| 47694 | 2531 | finally show ?thesis . | 
| 2532 | qed | |
| 2533 | ||
| 2534 | lemma measure_uniform_measure[simp]: | |
| 2535 | assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>" and B: "B \<in> sets M" | |
| 2536 | shows "measure (uniform_measure M A) B = measure M (A \<inter> B) / measure M A" | |
| 2537 | using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2538 | by (cases "emeasure M A" "emeasure M (A \<inter> B)" rule: ennreal2_cases) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2539 | (simp_all add: measure_def divide_ennreal top_ennreal.rep_eq top_ereal_def ennreal_top_divide) | 
| 47694 | 2540 | |
| 58606 | 2541 | lemma AE_uniform_measureI: | 
| 2542 | "A \<in> sets M \<Longrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x) \<Longrightarrow> (AE x in uniform_measure M A. P x)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2543 | unfolding uniform_measure_def by (auto simp: AE_density divide_ennreal_def) | 
| 58606 | 2544 | |
| 59000 | 2545 | lemma emeasure_uniform_measure_1: | 
| 2546 | "emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> emeasure (uniform_measure M S) S = 1" | |
| 2547 | by (subst emeasure_uniform_measure) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2548 | (simp_all add: emeasure_neq_0_sets emeasure_eq_ennreal_measure divide_ennreal | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2549 | zero_less_iff_neq_zero[symmetric]) | 
| 59000 | 2550 | |
| 2551 | lemma nn_integral_uniform_measure: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2552 | assumes f[measurable]: "f \<in> borel_measurable M" and S[measurable]: "S \<in> sets M" | 
| 59000 | 2553 | shows "(\<integral>\<^sup>+x. f x \<partial>uniform_measure M S) = (\<integral>\<^sup>+x. f x * indicator S x \<partial>M) / emeasure M S" | 
| 2554 | proof - | |
| 2555 |   { assume "emeasure M S = \<infinity>"
 | |
| 2556 | then have ?thesis | |
| 2557 | by (simp add: uniform_measure_def nn_integral_density f) } | |
| 2558 | moreover | |
| 2559 |   { assume [simp]: "emeasure M S = 0"
 | |
| 2560 | then have ae: "AE x in M. x \<notin> S" | |
| 2561 | using sets.sets_into_space[OF S] | |
| 2562 | by (subst AE_iff_measurable[OF _ refl]) (simp_all add: subset_eq cong: rev_conj_cong) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2563 | from ae have "(\<integral>\<^sup>+ x. indicator S x / 0 * f x \<partial>M) = 0" | 
| 59000 | 2564 | by (subst nn_integral_0_iff_AE) auto | 
| 2565 | moreover from ae have "(\<integral>\<^sup>+ x. f x * indicator S x \<partial>M) = 0" | |
| 2566 | by (subst nn_integral_0_iff_AE) auto | |
| 2567 | ultimately have ?thesis | |
| 2568 | by (simp add: uniform_measure_def nn_integral_density f) } | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2569 | moreover have "emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> ?thesis" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2570 | unfolding uniform_measure_def | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2571 | by (subst nn_integral_density) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2572 | (auto simp: ennreal_times_divide f nn_integral_divide[symmetric] mult.commute) | 
| 59000 | 2573 | ultimately show ?thesis by blast | 
| 2574 | qed | |
| 2575 | ||
| 2576 | lemma AE_uniform_measure: | |
| 2577 | assumes "emeasure M A \<noteq> 0" "emeasure M A < \<infinity>" | |
| 2578 | shows "(AE x in uniform_measure M A. P x) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x)" | |
| 2579 | proof - | |
| 2580 | have "A \<in> sets M" | |
| 61808 | 2581 | using \<open>emeasure M A \<noteq> 0\<close> by (metis emeasure_notin_sets) | 
| 59000 | 2582 | moreover have "\<And>x. 0 < indicator A x / emeasure M A \<longleftrightarrow> x \<in> A" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2583 | using assms | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2584 | by (cases "emeasure M A") (auto split: split_indicator simp: ennreal_zero_less_divide) | 
| 59000 | 2585 | ultimately show ?thesis | 
| 2586 | unfolding uniform_measure_def by (simp add: AE_density) | |
| 2587 | qed | |
| 2588 | ||
| 70136 | 2589 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Null measure\<close> | 
| 59425 | 2590 | |
| 2591 | lemma null_measure_eq_density: "null_measure M = density M (\<lambda>_. 0)" | |
| 2592 | by (intro measure_eqI) (simp_all add: emeasure_density) | |
| 2593 | ||
| 2594 | lemma nn_integral_null_measure[simp]: "(\<integral>\<^sup>+x. f x \<partial>null_measure M) = 0" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2595 | by (auto simp add: nn_integral_def simple_integral_def SUP_constant bot_ennreal_def le_fun_def | 
| 59425 | 2596 | intro!: exI[of _ "\<lambda>x. 0"]) | 
| 2597 | ||
| 2598 | lemma density_null_measure[simp]: "density (null_measure M) f = null_measure M" | |
| 2599 | proof (intro measure_eqI) | |
| 2600 | fix A show "emeasure (density (null_measure M) f) A = emeasure (null_measure M) A" | |
| 2601 | by (simp add: density_def) (simp only: null_measure_def[symmetric] emeasure_null_measure) | |
| 2602 | qed simp | |
| 2603 | ||
| 61808 | 2604 | subsubsection \<open>Uniform count measure\<close> | 
| 47694 | 2605 | |
| 70136 | 2606 | definition\<^marker>\<open>tag important\<close> "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2607 | |
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2608 | lemma | 
| 47694 | 2609 | shows space_uniform_count_measure: "space (uniform_count_measure A) = A" | 
| 2610 | and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" | |
| 2611 | unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) | |
| 59048 | 2612 | |
| 2613 | lemma sets_uniform_count_measure_count_space[measurable_cong]: | |
| 2614 | "sets (uniform_count_measure A) = sets (count_space A)" | |
| 2615 | by (simp add: sets_uniform_count_measure) | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2616 | |
| 47694 | 2617 | lemma emeasure_uniform_count_measure: | 
| 2618 | "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (uniform_count_measure A) X = card X / card A" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2619 | by (simp add: emeasure_point_measure_finite uniform_count_measure_def divide_inverse ennreal_mult | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2620 | ennreal_of_nat_eq_real_of_nat) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2621 | |
| 79583 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2622 | lemma emeasure_uniform_count_measure_if: | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2623 | "finite A \<Longrightarrow> emeasure (uniform_count_measure A) X = (if X \<subseteq> A then card X / card A else 0)" | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2624 | by (simp add: emeasure_notin_sets emeasure_uniform_count_measure sets_uniform_count_measure) | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2625 | |
| 47694 | 2626 | lemma measure_uniform_count_measure: | 
| 2627 | "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> measure (uniform_count_measure A) X = card X / card A" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2628 | by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def enn2real_mult) | 
| 47694 | 2629 | |
| 79583 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2630 | lemma measure_uniform_count_measure_if: | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2631 | "finite A \<Longrightarrow> measure (uniform_count_measure A) X = (if X \<subseteq> A then card X / card A else 0)" | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2632 | by (simp add: measure_uniform_count_measure measure_notin_sets sets_uniform_count_measure) | 
| 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2633 | |
| 61633 | 2634 | lemma space_uniform_count_measure_empty_iff [simp]: | 
| 2635 |   "space (uniform_count_measure X) = {} \<longleftrightarrow> X = {}"
 | |
| 2636 | by(simp add: space_uniform_count_measure) | |
| 2637 | ||
| 2638 | lemma sets_uniform_count_measure_eq_UNIV [simp]: | |
| 2639 | "sets (uniform_count_measure UNIV) = UNIV \<longleftrightarrow> True" | |
| 2640 | "UNIV = sets (uniform_count_measure UNIV) \<longleftrightarrow> True" | |
| 79583 
a521c241e946
Further lemmas concerning complexity and measures
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 2641 | by(simp_all add: sets_uniform_count_measure) | 
| 61633 | 2642 | |
| 70136 | 2643 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Scaled measure\<close> | 
| 61634 | 2644 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2645 | lemma nn_integral_scale_measure: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2646 | assumes f: "f \<in> borel_measurable M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2647 | shows "nn_integral (scale_measure r M) f = r * nn_integral M f" | 
| 61634 | 2648 | using f | 
| 2649 | proof induction | |
| 2650 | case (cong f g) | |
| 2651 | thus ?case | |
| 2652 | by(simp add: cong.hyps space_scale_measure cong: nn_integral_cong_simp) | |
| 2653 | next | |
| 2654 | case (mult f c) | |
| 2655 | thus ?case | |
| 2656 | by(simp add: nn_integral_cmult max_def mult.assoc mult.left_commute) | |
| 2657 | next | |
| 2658 | case (add f g) | |
| 2659 | thus ?case | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2660 | by(simp add: nn_integral_add distrib_left) | 
| 61634 | 2661 | next | 
| 2662 | case (seq U) | |
| 2663 | thus ?case | |
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 2664 | by(simp add: nn_integral_monotone_convergence_SUP SUP_mult_left_ennreal image_comp) | 
| 61634 | 2665 | qed simp | 
| 2666 | ||
| 35748 | 2667 | end |