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