src/HOL/Probability/Lebesgue_Measure.thy
author paulson <lp15@cam.ac.uk>
Wed, 01 Apr 2015 16:04:21 +0100
changeset 59872 db4000b71fdb
parent 59741 5b762cd73a8e
child 60615 e5fa1d5d3952
permissions -rw-r--r--
Theorem "arctan" is no longer a default simprule
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     1
(*  Title:      HOL/Probability/Lebesgue_Measure.thy
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:     Robert Himmelmann, TU München
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
     4
    Author:     Jeremy Avigad
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
     5
    Author:     Luke Serafin
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     6
*)
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     7
59357
f366643536cd allow line breaks in integral notation
Andreas Lochbihler
parents: 59048
diff changeset
     8
section {* Lebesgue measure *}
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     9
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    10
theory Lebesgue_Measure
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    11
  imports Finite_Product_Measure Bochner_Integration Caratheodory
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    12
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    13
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    14
subsection {* Every right continuous and nondecreasing function gives rise to a measure *}
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    15
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    16
definition interval_measure :: "(real \<Rightarrow> real) \<Rightarrow> real measure" where
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    17
  "interval_measure F = extend_measure UNIV {(a, b). a \<le> b} (\<lambda>(a, b). {a <.. b}) (\<lambda>(a, b). ereal (F b - F a))"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    18
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    19
lemma emeasure_interval_measure_Ioc:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    20
  assumes "a \<le> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    21
  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    22
  assumes right_cont_F : "\<And>a. continuous (at_right a) F" 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    23
  shows "emeasure (interval_measure F) {a <.. b} = F b - F a"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    24
proof (rule extend_measure_caratheodory_pair[OF interval_measure_def `a \<le> b`])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    25
  show "semiring_of_sets UNIV {{a<..b} |a b :: real. a \<le> b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    26
  proof (unfold_locales, safe)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    27
    fix a b c d :: real assume *: "a \<le> b" "c \<le> d"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    28
    then show "\<exists>C\<subseteq>{{a<..b} |a b. a \<le> b}. finite C \<and> disjoint C \<and> {a<..b} - {c<..d} = \<Union>C"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    29
    proof cases
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    30
      let ?C = "{{a<..b}}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    31
      assume "b < c \<or> d \<le> a \<or> d \<le> c"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    32
      with * have "?C \<subseteq> {{a<..b} |a b. a \<le> b} \<and> finite ?C \<and> disjoint ?C \<and> {a<..b} - {c<..d} = \<Union>?C"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    33
        by (auto simp add: disjoint_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    34
      thus ?thesis ..
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    35
    next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    36
      let ?C = "{{a<..c}, {d<..b}}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    37
      assume "\<not> (b < c \<or> d \<le> a \<or> d \<le> c)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    38
      with * have "?C \<subseteq> {{a<..b} |a b. a \<le> b} \<and> finite ?C \<and> disjoint ?C \<and> {a<..b} - {c<..d} = \<Union>?C"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    39
        by (auto simp add: disjoint_def Ioc_inj) (metis linear)+
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    40
      thus ?thesis ..
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    41
    qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    42
  qed (auto simp: Ioc_inj, metis linear)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    43
  
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    44
next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    45
  fix l r :: "nat \<Rightarrow> real" and a b :: real
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    46
  assume l_r[simp]: "\<And>n. l n \<le> r n" and "a \<le> b" and disj: "disjoint_family (\<lambda>n. {l n<..r n})" 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    47
  assume lr_eq_ab: "(\<Union>i. {l i<..r i}) = {a<..b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    48
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    49
  have [intro, simp]: "\<And>a b. a \<le> b \<Longrightarrow> 0 \<le> F b - F a"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    50
    by (auto intro!: l_r mono_F simp: diff_le_iff)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    51
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    52
  { fix S :: "nat set" assume "finite S"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    53
    moreover note `a \<le> b`
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    54
    moreover have "\<And>i. i \<in> S \<Longrightarrow> {l i <.. r i} \<subseteq> {a <.. b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    55
      unfolding lr_eq_ab[symmetric] by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    56
    ultimately have "(\<Sum>i\<in>S. F (r i) - F (l i)) \<le> F b - F a"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    57
    proof (induction S arbitrary: a rule: finite_psubset_induct)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    58
      case (psubset S)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    59
      show ?case
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    60
      proof cases
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    61
        assume "\<exists>i\<in>S. l i < r i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    62
        with `finite S` have "Min (l ` {i\<in>S. l i < r i}) \<in> l ` {i\<in>S. l i < r i}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    63
          by (intro Min_in) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    64
        then obtain m where m: "m \<in> S" "l m < r m" "l m = Min (l ` {i\<in>S. l i < r i})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    65
          by fastforce
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    66
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    67
        have "(\<Sum>i\<in>S. F (r i) - F (l i)) = (F (r m) - F (l m)) + (\<Sum>i\<in>S - {m}. F (r i) - F (l i))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    68
          using m psubset by (intro setsum.remove) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    69
        also have "(\<Sum>i\<in>S - {m}. F (r i) - F (l i)) \<le> F b - F (r m)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    70
        proof (intro psubset.IH)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    71
          show "S - {m} \<subset> S"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    72
            using `m\<in>S` by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    73
          show "r m \<le> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    74
            using psubset.prems(2)[OF `m\<in>S`] `l m < r m` by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    75
        next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    76
          fix i assume "i \<in> S - {m}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    77
          then have i: "i \<in> S" "i \<noteq> m" by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    78
          { assume i': "l i < r i" "l i < r m"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    79
            moreover with `finite S` i m have "l m \<le> l i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    80
              by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    81
            ultimately have "{l i <.. r i} \<inter> {l m <.. r m} \<noteq> {}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    82
              by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    83
            then have False
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    84
              using disjoint_family_onD[OF disj, of i m] i by auto }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    85
          then have "l i \<noteq> r i \<Longrightarrow> r m \<le> l i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    86
            unfolding not_less[symmetric] using l_r[of i] by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    87
          then show "{l i <.. r i} \<subseteq> {r m <.. b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    88
            using psubset.prems(2)[OF `i\<in>S`] by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    89
        qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    90
        also have "F (r m) - F (l m) \<le> F (r m) - F a"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    91
          using psubset.prems(2)[OF `m \<in> S`] `l m < r m`
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    92
          by (auto simp add: Ioc_subset_iff intro!: mono_F)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    93
        finally show ?case
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    94
          by (auto intro: add_mono)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    95
      qed (simp add: `a \<le> b` less_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    96
    qed }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    97
  note claim1 = this
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    98
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    99
  (* second key induction: a lower bound on the measures of any finite collection of Ai's
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   100
     that cover an interval {u..v} *)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   101
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   102
  { fix S u v and l r :: "nat \<Rightarrow> real"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   103
    assume "finite S" "\<And>i. i\<in>S \<Longrightarrow> l i < r i" "{u..v} \<subseteq> (\<Union>i\<in>S. {l i<..< r i})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   104
    then have "F v - F u \<le> (\<Sum>i\<in>S. F (r i) - F (l i))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   105
    proof (induction arbitrary: v u rule: finite_psubset_induct)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   106
      case (psubset S)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   107
      show ?case
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   108
      proof cases
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   109
        assume "S = {}" then show ?case
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   110
          using psubset by (simp add: mono_F)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   111
      next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   112
        assume "S \<noteq> {}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   113
        then obtain j where "j \<in> S"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   114
          by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   115
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   116
        let ?R = "r j < u \<or> l j > v \<or> (\<exists>i\<in>S-{j}. l i \<le> l j \<and> r j \<le> r i)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   117
        show ?case
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   118
        proof cases
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   119
          assume "?R"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   120
          with `j \<in> S` psubset.prems have "{u..v} \<subseteq> (\<Union>i\<in>S-{j}. {l i<..< r i})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   121
            apply (auto simp: subset_eq Ball_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   122
            apply (metis Diff_iff less_le_trans leD linear singletonD)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   123
            apply (metis Diff_iff less_le_trans leD linear singletonD)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   124
            apply (metis order_trans less_le_not_le linear)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   125
            done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   126
          with `j \<in> S` have "F v - F u \<le> (\<Sum>i\<in>S - {j}. F (r i) - F (l i))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   127
            by (intro psubset) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   128
          also have "\<dots> \<le> (\<Sum>i\<in>S. F (r i) - F (l i))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   129
            using psubset.prems
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   130
            by (intro setsum_mono2 psubset) (auto intro: less_imp_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   131
          finally show ?thesis .
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   132
        next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   133
          assume "\<not> ?R"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   134
          then have j: "u \<le> r j" "l j \<le> v" "\<And>i. i \<in> S - {j} \<Longrightarrow> r i < r j \<or> l i > l j"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   135
            by (auto simp: not_less)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   136
          let ?S1 = "{i \<in> S. l i < l j}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   137
          let ?S2 = "{i \<in> S. r i > r j}"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   138
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   139
          have "(\<Sum>i\<in>S. F (r i) - F (l i)) \<ge> (\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   140
            using `j \<in> S` `finite S` psubset.prems j
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   141
            by (intro setsum_mono2) (auto intro: less_imp_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   142
          also have "(\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i)) =
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   143
            (\<Sum>i\<in>?S1. F (r i) - F (l i)) + (\<Sum>i\<in>?S2 . F (r i) - F (l i)) + (F (r j) - F (l j))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   144
            using psubset(1) psubset.prems(1) j
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   145
            apply (subst setsum.union_disjoint)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   146
            apply simp_all
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   147
            apply (subst setsum.union_disjoint)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   148
            apply auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   149
            apply (metis less_le_not_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   150
            done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   151
          also (xtrans) have "(\<Sum>i\<in>?S1. F (r i) - F (l i)) \<ge> F (l j) - F u"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   152
            using `j \<in> S` `finite S` psubset.prems j
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   153
            apply (intro psubset.IH psubset)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   154
            apply (auto simp: subset_eq Ball_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   155
            apply (metis less_le_trans not_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   156
            done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   157
          also (xtrans) have "(\<Sum>i\<in>?S2. F (r i) - F (l i)) \<ge> F v - F (r j)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   158
            using `j \<in> S` `finite S` psubset.prems j
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   159
            apply (intro psubset.IH psubset)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   160
            apply (auto simp: subset_eq Ball_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   161
            apply (metis le_less_trans not_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   162
            done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   163
          finally (xtrans) show ?case
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   164
            by (auto simp: add_mono)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   165
        qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   166
      qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   167
    qed }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   168
  note claim2 = this
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   169
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   170
  (* now prove the inequality going the other way *)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   171
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   172
  { fix epsilon :: real assume egt0: "epsilon > 0"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   173
    have "\<forall>i. \<exists>d. d > 0 &  F (r i + d) < F (r i) + epsilon / 2^(i+2)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   174
    proof 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   175
      fix i
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   176
      note right_cont_F [of "r i"]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   177
      thus "\<exists>d. d > 0 \<and> F (r i + d) < F (r i) + epsilon / 2^(i+2)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   178
        apply -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   179
        apply (subst (asm) continuous_at_right_real_increasing)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   180
        apply (rule mono_F, assumption)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   181
        apply (drule_tac x = "epsilon / 2 ^ (i + 2)" in spec)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   182
        apply (erule impE)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   183
        using egt0 by (auto simp add: field_simps)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   184
    qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   185
    then obtain delta where 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   186
        deltai_gt0: "\<And>i. delta i > 0" and
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   187
        deltai_prop: "\<And>i. F (r i + delta i) < F (r i) + epsilon / 2^(i+2)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   188
      by metis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   189
    have "\<exists>a' > a. F a' - F a < epsilon / 2"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   190
      apply (insert right_cont_F [of a])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   191
      apply (subst (asm) continuous_at_right_real_increasing)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   192
      using mono_F apply force
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   193
      apply (drule_tac x = "epsilon / 2" in spec)
59554
4044f53326c9 inlined rules to free user-space from technical names
haftmann
parents: 59425
diff changeset
   194
      using egt0 unfolding mult.commute [of 2] by force
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   195
    then obtain a' where a'lea [arith]: "a' > a" and 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   196
      a_prop: "F a' - F a < epsilon / 2"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   197
      by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   198
    def S' \<equiv> "{i. l i < r i}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   199
    obtain S :: "nat set" where 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   200
      "S \<subseteq> S'" and finS: "finite S" and 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   201
      Sprop: "{a'..b} \<subseteq> (\<Union>i \<in> S. {l i<..<r i + delta i})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   202
    proof (rule compactE_image)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   203
      show "compact {a'..b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   204
        by (rule compact_Icc)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   205
      show "\<forall>i \<in> S'. open ({l i<..<r i + delta i})" by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   206
      have "{a'..b} \<subseteq> {a <.. b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   207
        by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   208
      also have "{a <.. b} = (\<Union>i\<in>S'. {l i<..r i})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   209
        unfolding lr_eq_ab[symmetric] by (fastforce simp add: S'_def intro: less_le_trans)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   210
      also have "\<dots> \<subseteq> (\<Union>i \<in> S'. {l i<..<r i + delta i})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   211
        apply (intro UN_mono)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   212
        apply (auto simp: S'_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   213
        apply (cut_tac i=i in deltai_gt0)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   214
        apply simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   215
        done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   216
      finally show "{a'..b} \<subseteq> (\<Union>i \<in> S'. {l i<..<r i + delta i})" .
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   217
    qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   218
    with S'_def have Sprop2: "\<And>i. i \<in> S \<Longrightarrow> l i < r i" by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   219
    from finS have "\<exists>n. \<forall>i \<in> S. i \<le> n" 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   220
      by (subst finite_nat_set_iff_bounded_le [symmetric])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   221
    then obtain n where Sbound [rule_format]: "\<forall>i \<in> S. i \<le> n" ..
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   222
    have "F b - F a' \<le> (\<Sum>i\<in>S. F (r i + delta i) - F (l i))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   223
      apply (rule claim2 [rule_format])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   224
      using finS Sprop apply auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   225
      apply (frule Sprop2)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   226
      apply (subgoal_tac "delta i > 0")
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   227
      apply arith
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   228
      by (rule deltai_gt0)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   229
    also have "... \<le> (SUM i : S. F(r i) - F(l i) + epsilon / 2^(i+2))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   230
      apply (rule setsum_mono)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   231
      apply simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   232
      apply (rule order_trans)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   233
      apply (rule less_imp_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   234
      apply (rule deltai_prop)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   235
      by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   236
    also have "... = (SUM i : S. F(r i) - F(l i)) + 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   237
        (epsilon / 4) * (SUM i : S. (1 / 2)^i)" (is "_ = ?t + _")
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   238
      by (subst setsum.distrib) (simp add: field_simps setsum_right_distrib)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   239
    also have "... \<le> ?t + (epsilon / 4) * (\<Sum> i < Suc n. (1 / 2)^i)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   240
      apply (rule add_left_mono)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   241
      apply (rule mult_left_mono)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   242
      apply (rule setsum_mono2)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   243
      using egt0 apply auto 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   244
      by (frule Sbound, auto)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   245
    also have "... \<le> ?t + (epsilon / 2)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   246
      apply (rule add_left_mono)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   247
      apply (subst geometric_sum)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   248
      apply auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   249
      apply (rule mult_left_mono)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   250
      using egt0 apply auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   251
      done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   252
    finally have aux2: "F b - F a' \<le> (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon / 2"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   253
      by simp
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   254
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   255
    have "F b - F a = (F b - F a') + (F a' - F a)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   256
      by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   257
    also have "... \<le> (F b - F a') + epsilon / 2"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   258
      using a_prop by (intro add_left_mono) simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   259
    also have "... \<le> (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon / 2 + epsilon / 2"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   260
      apply (intro add_right_mono)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   261
      apply (rule aux2)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   262
      done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   263
    also have "... = (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   264
      by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   265
    also have "... \<le> (\<Sum>i\<le>n. F (r i) - F (l i)) + epsilon"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   266
      using finS Sbound Sprop by (auto intro!: add_right_mono setsum_mono3)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   267
    finally have "ereal (F b - F a) \<le> (\<Sum>i\<le>n. ereal (F (r i) - F (l i))) + epsilon"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   268
      by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   269
    then have "ereal (F b - F a) \<le> (\<Sum>i. ereal (F (r i) - F (l i))) + (epsilon :: real)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   270
      apply (rule_tac order_trans)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   271
      prefer 2
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   272
      apply (rule add_mono[where c="ereal epsilon"])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   273
      apply (rule suminf_upper[of _ "Suc n"])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   274
      apply (auto simp add: lessThan_Suc_atMost)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   275
      done }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   276
  hence "ereal (F b - F a) \<le> (\<Sum>i. ereal (F (r i) - F (l i)))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   277
    by (auto intro: ereal_le_epsilon2)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   278
  moreover
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   279
  have "(\<Sum>i. ereal (F (r i) - F (l i))) \<le> ereal (F b - F a)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   280
    by (auto simp add: claim1 intro!: suminf_bound)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   281
  ultimately show "(\<Sum>n. ereal (F (r n) - F (l n))) = ereal (F b - F a)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   282
    by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   283
qed (auto simp: Ioc_inj diff_le_iff mono_F)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   284
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   285
lemma measure_interval_measure_Ioc:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   286
  assumes "a \<le> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   287
  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   288
  assumes right_cont_F : "\<And>a. continuous (at_right a) F" 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   289
  shows "measure (interval_measure F) {a <.. b} = F b - F a"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   290
  unfolding measure_def
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   291
  apply (subst emeasure_interval_measure_Ioc)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   292
  apply fact+
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   293
  apply simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   294
  done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   295
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   296
lemma emeasure_interval_measure_Ioc_eq:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   297
  "(\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y) \<Longrightarrow> (\<And>a. continuous (at_right a) F) \<Longrightarrow>
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   298
    emeasure (interval_measure F) {a <.. b} = (if a \<le> b then F b - F a else 0)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   299
  using emeasure_interval_measure_Ioc[of a b F] by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   300
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59000
diff changeset
   301
lemma sets_interval_measure [simp, measurable_cong]: "sets (interval_measure F) = sets borel"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   302
  apply (simp add: sets_extend_measure interval_measure_def borel_sigma_sets_Ioc)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   303
  apply (rule sigma_sets_eqI)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   304
  apply auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   305
  apply (case_tac "a \<le> ba")
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   306
  apply (auto intro: sigma_sets.Empty)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   307
  done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   308
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   309
lemma space_interval_measure [simp]: "space (interval_measure F) = UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   310
  by (simp add: interval_measure_def space_extend_measure)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   311
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   312
lemma emeasure_interval_measure_Icc:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   313
  assumes "a \<le> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   314
  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   315
  assumes cont_F : "continuous_on UNIV F" 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   316
  shows "emeasure (interval_measure F) {a .. b} = F b - F a"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   317
proof (rule tendsto_unique)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   318
  { fix a b :: real assume "a \<le> b" then have "emeasure (interval_measure F) {a <.. b} = F b - F a"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   319
      using cont_F
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   320
      by (subst emeasure_interval_measure_Ioc)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   321
         (auto intro: mono_F continuous_within_subset simp: continuous_on_eq_continuous_within) }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   322
  note * = this
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   323
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   324
  let ?F = "interval_measure F"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   325
  show "((\<lambda>a. F b - F a) ---> emeasure ?F {a..b}) (at_left a)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   326
  proof (rule tendsto_at_left_sequentially)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   327
    show "a - 1 < a" by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   328
    fix X assume "\<And>n. X n < a" "incseq X" "X ----> a"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   329
    with `a \<le> b` have "(\<lambda>n. emeasure ?F {X n<..b}) ----> emeasure ?F (\<Inter>n. {X n <..b})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   330
      apply (intro Lim_emeasure_decseq)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   331
      apply (auto simp: decseq_def incseq_def emeasure_interval_measure_Ioc *)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   332
      apply force
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   333
      apply (subst (asm ) *)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   334
      apply (auto intro: less_le_trans less_imp_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   335
      done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   336
    also have "(\<Inter>n. {X n <..b}) = {a..b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   337
      using `\<And>n. X n < a`
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   338
      apply auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   339
      apply (rule LIMSEQ_le_const2[OF `X ----> a`])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   340
      apply (auto intro: less_imp_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   341
      apply (auto intro: less_le_trans)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   342
      done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   343
    also have "(\<lambda>n. emeasure ?F {X n<..b}) = (\<lambda>n. F b - F (X n))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   344
      using `\<And>n. X n < a` `a \<le> b` by (subst *) (auto intro: less_imp_le less_le_trans)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   345
    finally show "(\<lambda>n. F b - F (X n)) ----> emeasure ?F {a..b}" .
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   346
  qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   347
  show "((\<lambda>a. ereal (F b - F a)) ---> F b - F a) (at_left a)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   348
    using cont_F
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   349
    by (intro lim_ereal[THEN iffD2] tendsto_intros )
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   350
       (auto simp: continuous_on_def intro: tendsto_within_subset)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   351
qed (rule trivial_limit_at_left_real)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   352
  
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   353
lemma sigma_finite_interval_measure:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   354
  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   355
  assumes right_cont_F : "\<And>a. continuous (at_right a) F" 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   356
  shows "sigma_finite_measure (interval_measure F)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   357
  apply unfold_locales
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   358
  apply (intro exI[of _ "(\<lambda>(a, b). {a <.. b}) ` (\<rat> \<times> \<rat>)"])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   359
  apply (auto intro!: Rats_no_top_le Rats_no_bot_less countable_rat simp: emeasure_interval_measure_Ioc_eq[OF assms])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   360
  done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   361
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   362
subsection {* Lebesgue-Borel measure *}
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   363
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   364
definition lborel :: "('a :: euclidean_space) measure" where
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   365
  "lborel = distr (\<Pi>\<^sub>M b\<in>Basis. interval_measure (\<lambda>x. x)) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   366
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   367
lemma 
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59000
diff changeset
   368
  shows sets_lborel[simp, measurable_cong]: "sets lborel = sets borel"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   369
    and space_lborel[simp]: "space lborel = space borel"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   370
    and measurable_lborel1[simp]: "measurable M lborel = measurable M borel"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   371
    and measurable_lborel2[simp]: "measurable lborel M = measurable borel M"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   372
  by (simp_all add: lborel_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   373
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   374
context
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   375
begin
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   376
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   377
interpretation sigma_finite_measure "interval_measure (\<lambda>x. x)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   378
  by (rule sigma_finite_interval_measure) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   379
interpretation finite_product_sigma_finite "\<lambda>_. interval_measure (\<lambda>x. x)" Basis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   380
  proof qed simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   381
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   382
lemma lborel_eq_real: "lborel = interval_measure (\<lambda>x. x)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   383
  unfolding lborel_def Basis_real_def
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   384
  using distr_id[of "interval_measure (\<lambda>x. x)"]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   385
  by (subst distr_component[symmetric])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   386
     (simp_all add: distr_distr comp_def del: distr_id cong: distr_cong)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   387
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   388
lemma lborel_eq: "lborel = distr (\<Pi>\<^sub>M b\<in>Basis. lborel) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   389
  by (subst lborel_def) (simp add: lborel_eq_real)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   390
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   391
lemma nn_integral_lborel_setprod:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   392
  assumes [measurable]: "\<And>b. b \<in> Basis \<Longrightarrow> f b \<in> borel_measurable borel"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   393
  assumes nn[simp]: "\<And>b x. b \<in> Basis \<Longrightarrow> 0 \<le> f b x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   394
  shows "(\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. f b (x \<bullet> b)) \<partial>lborel) = (\<Prod>b\<in>Basis. (\<integral>\<^sup>+x. f b x \<partial>lborel))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   395
  by (simp add: lborel_def nn_integral_distr product_nn_integral_setprod
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   396
                product_nn_integral_singleton)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   397
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   398
lemma emeasure_lborel_Icc[simp]: 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   399
  fixes l u :: real
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   400
  assumes [simp]: "l \<le> u"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   401
  shows "emeasure lborel {l .. u} = u - l"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   402
proof -
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   403
  have "((\<lambda>f. f 1) -` {l..u} \<inter> space (Pi\<^sub>M {1} (\<lambda>b. interval_measure (\<lambda>x. x)))) = {1::real} \<rightarrow>\<^sub>E {l..u}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   404
    by (auto simp: space_PiM)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   405
  then show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   406
    by (simp add: lborel_def emeasure_distr emeasure_PiM emeasure_interval_measure_Icc continuous_on_id)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   407
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   408
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   409
lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ereal (if l \<le> u then u - l else 0)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   410
  by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   411
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   412
lemma emeasure_lborel_cbox[simp]:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   413
  assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   414
  shows "emeasure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   415
proof -
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   416
  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b) :: ereal) = indicator (cbox l u)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   417
    by (auto simp: fun_eq_iff cbox_def setprod_ereal_0 split: split_indicator)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   418
  then have "emeasure lborel (cbox l u) = (\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b)) \<partial>lborel)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   419
    by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   420
  also have "\<dots> = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   421
    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ereal inner_diff_left)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   422
  finally show ?thesis .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   423
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   424
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   425
lemma AE_lborel_singleton: "AE x in lborel::'a::euclidean_space measure. x \<noteq> c"
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 57514
diff changeset
   426
  using SOME_Basis AE_discrete_difference [of "{c}" lborel]
af9eb5e566dd eliminated redundancies;
haftmann
parents: 57514
diff changeset
   427
    emeasure_lborel_cbox [of c c] by (auto simp add: cbox_sing)
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   428
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   429
lemma emeasure_lborel_Ioo[simp]:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   430
  assumes [simp]: "l \<le> u"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   431
  shows "emeasure lborel {l <..< u} = ereal (u - l)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   432
proof -
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   433
  have "emeasure lborel {l <..< u} = emeasure lborel {l .. u}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   434
    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   435
  then show ?thesis
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   436
    by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   437
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   438
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   439
lemma emeasure_lborel_Ioc[simp]:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   440
  assumes [simp]: "l \<le> u"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   441
  shows "emeasure lborel {l <.. u} = ereal (u - l)"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   442
proof -
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   443
  have "emeasure lborel {l <.. u} = emeasure lborel {l .. u}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   444
    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   445
  then show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   446
    by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   447
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   448
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   449
lemma emeasure_lborel_Ico[simp]:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   450
  assumes [simp]: "l \<le> u"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   451
  shows "emeasure lborel {l ..< u} = ereal (u - l)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   452
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   453
  have "emeasure lborel {l ..< u} = emeasure lborel {l .. u}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   454
    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   455
  then show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   456
    by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   457
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   458
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   459
lemma emeasure_lborel_box[simp]:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   460
  assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   461
  shows "emeasure lborel (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   462
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   463
  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b <..< u\<bullet>b} (x \<bullet> b) :: ereal) = indicator (box l u)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   464
    by (auto simp: fun_eq_iff box_def setprod_ereal_0 split: split_indicator)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   465
  then have "emeasure lborel (box l u) = (\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. indicator {l\<bullet>b <..< u\<bullet>b} (x \<bullet> b)) \<partial>lborel)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   466
    by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   467
  also have "\<dots> = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   468
    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ereal inner_diff_left)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   469
  finally show ?thesis .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   470
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   471
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   472
lemma emeasure_lborel_cbox_eq:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   473
  "emeasure lborel (cbox l u) = (if \<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b then \<Prod>b\<in>Basis. (u - l) \<bullet> b else 0)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   474
  using box_eq_empty(2)[THEN iffD2, of u l] by (auto simp: not_le)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   475
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   476
lemma emeasure_lborel_box_eq:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   477
  "emeasure lborel (box l u) = (if \<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b then \<Prod>b\<in>Basis. (u - l) \<bullet> b else 0)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   478
  using box_eq_empty(1)[THEN iffD2, of u l] by (auto simp: not_le dest!: less_imp_le) force
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   479
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   480
lemma
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   481
  fixes l u :: real
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   482
  assumes [simp]: "l \<le> u"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   483
  shows measure_lborel_Icc[simp]: "measure lborel {l .. u} = u - l"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   484
    and measure_lborel_Ico[simp]: "measure lborel {l ..< u} = u - l"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   485
    and measure_lborel_Ioc[simp]: "measure lborel {l <.. u} = u - l"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   486
    and measure_lborel_Ioo[simp]: "measure lborel {l <..< u} = u - l"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   487
  by (simp_all add: measure_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   488
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   489
lemma 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   490
  assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   491
  shows measure_lborel_box[simp]: "measure lborel (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   492
    and measure_lborel_cbox[simp]: "measure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   493
  by (simp_all add: measure_def)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   494
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   495
lemma sigma_finite_lborel: "sigma_finite_measure lborel"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   496
proof
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   497
  show "\<exists>A::'a set set. countable A \<and> A \<subseteq> sets lborel \<and> \<Union>A = space lborel \<and> (\<forall>a\<in>A. emeasure lborel a \<noteq> \<infinity>)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   498
    by (intro exI[of _ "range (\<lambda>n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One))"])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   499
       (auto simp: emeasure_lborel_cbox_eq UN_box_eq_UNIV)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   500
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   501
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   502
end
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
   503
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   504
lemma emeasure_lborel_UNIV: "emeasure lborel (UNIV::'a::euclidean_space set) = \<infinity>"
59741
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   505
proof -
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   506
  { fix n::nat
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   507
    let ?Ba = "Basis :: 'a set"
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   508
    have "real n \<le> (2::real) ^ card ?Ba * real n"
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   509
      by (simp add: mult_le_cancel_right1)
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   510
    also 
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   511
    have "... \<le> (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba"
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   512
      apply (rule mult_left_mono)
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   513
      apply (metis DIM_positive One_nat_def less_eq_Suc_le less_imp_le real_of_nat_le_iff real_of_nat_power self_le_power zero_less_Suc)
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   514
      apply (simp add: DIM_positive)
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   515
      done
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   516
    finally have "real n \<le> (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba" .
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   517
  } note [intro!] = this
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   518
  show ?thesis
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   519
    unfolding UN_box_eq_UNIV[symmetric]
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   520
    apply (subst SUP_emeasure_incseq[symmetric])
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   521
    apply (auto simp: incseq_def subset_box inner_add_left setprod_constant 
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   522
               intro!: SUP_PInfty)
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   523
    done 
5b762cd73a8e Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
paulson <lp15@cam.ac.uk>
parents: 59554
diff changeset
   524
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   525
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   526
lemma emeasure_lborel_singleton[simp]: "emeasure lborel {x} = 0"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   527
  using emeasure_lborel_cbox[of x x] nonempty_Basis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   528
  by (auto simp del: emeasure_lborel_cbox nonempty_Basis simp add: cbox_sing)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   529
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   530
lemma emeasure_lborel_countable:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   531
  fixes A :: "'a::euclidean_space set"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   532
  assumes "countable A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   533
  shows "emeasure lborel A = 0"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   534
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   535
  have "A \<subseteq> (\<Union>i. {from_nat_into A i})" using from_nat_into_surj assms by force
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   536
  moreover have "emeasure lborel (\<Union>i. {from_nat_into A i}) = 0"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   537
    by (rule emeasure_UN_eq_0) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   538
  ultimately have "emeasure lborel A \<le> 0" using emeasure_mono
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   539
    by (metis assms bot.extremum_unique emeasure_empty image_eq_UN range_from_nat_into sets.empty_sets)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   540
  thus ?thesis by (auto simp add: emeasure_le_0_iff)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   541
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   542
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   543
lemma countable_imp_null_set_lborel: "countable A \<Longrightarrow> A \<in> null_sets lborel"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   544
  by (simp add: null_sets_def emeasure_lborel_countable sets.countable)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   545
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   546
lemma finite_imp_null_set_lborel: "finite A \<Longrightarrow> A \<in> null_sets lborel"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   547
  by (intro countable_imp_null_set_lborel countable_finite)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   548
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   549
lemma lborel_neq_count_space[simp]: "lborel \<noteq> count_space (A::('a::ordered_euclidean_space) set)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   550
proof
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   551
  assume asm: "lborel = count_space A"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   552
  have "space lborel = UNIV" by simp
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   553
  hence [simp]: "A = UNIV" by (subst (asm) asm) (simp only: space_count_space)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   554
  have "emeasure lborel {undefined::'a} = 1" 
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   555
      by (subst asm, subst emeasure_count_space_finite) auto
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   556
  moreover have "emeasure lborel {undefined} \<noteq> 1" by simp
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   557
  ultimately show False by contradiction
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   558
qed
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   559
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   560
subsection {* Affine transformation on the Lebesgue-Borel *}
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   561
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   562
lemma lborel_eqI:
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   563
  fixes M :: "'a::euclidean_space measure"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   564
  assumes emeasure_eq: "\<And>l u. (\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b) \<Longrightarrow> emeasure M (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   565
  assumes sets_eq: "sets M = sets borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   566
  shows "lborel = M"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   567
proof (rule measure_eqI_generator_eq)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   568
  let ?E = "range (\<lambda>(a, b). box a b::'a set)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   569
  show "Int_stable ?E"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   570
    by (auto simp: Int_stable_def box_Int_box)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   571
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   572
  show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   573
    by (simp_all add: borel_eq_box sets_eq)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   574
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   575
  let ?A = "\<lambda>n::nat. box (- (real n *\<^sub>R One)) (real n *\<^sub>R One) :: 'a set"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   576
  show "range ?A \<subseteq> ?E" "(\<Union>i. ?A i) = UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   577
    unfolding UN_box_eq_UNIV by auto
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   578
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   579
  { fix i show "emeasure lborel (?A i) \<noteq> \<infinity>" by auto }
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   580
  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   581
      apply (auto simp: emeasure_eq emeasure_lborel_box_eq )
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   582
      apply (subst box_eq_empty(1)[THEN iffD2])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   583
      apply (auto intro: less_imp_le simp: not_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   584
      done }
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   585
qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   586
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   587
lemma lborel_affine:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   588
  fixes t :: "'a::euclidean_space" assumes "c \<noteq> 0"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   589
  shows "lborel = density (distr lborel borel (\<lambda>x. t + c *\<^sub>R x)) (\<lambda>_. \<bar>c\<bar>^DIM('a))" (is "_ = ?D")
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   590
proof (rule lborel_eqI)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   591
  let ?B = "Basis :: 'a set"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   592
  fix l u assume le: "\<And>b. b \<in> ?B \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   593
  show "emeasure ?D (box l u) = (\<Prod>b\<in>?B. (u - l) \<bullet> b)"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   594
  proof cases
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   595
    assume "0 < c"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   596
    then have "(\<lambda>x. t + c *\<^sub>R x) -` box l u = box ((l - t) /\<^sub>R c) ((u - t) /\<^sub>R c)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   597
      by (auto simp: field_simps box_def inner_simps)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   598
    with `0 < c` show ?thesis
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   599
      using le
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   600
      by (auto simp: field_simps emeasure_density nn_integral_distr nn_integral_cmult
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   601
                     emeasure_lborel_box_eq inner_simps setprod_dividef setprod_constant
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   602
                     borel_measurable_indicator' emeasure_distr)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   603
  next
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   604
    assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   605
    then have "box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c) = (\<lambda>x. t + c *\<^sub>R x) -` box l u"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   606
      by (auto simp: field_simps box_def inner_simps)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   607
    then have "\<And>x. indicator (box l u) (t + c *\<^sub>R x) = (indicator (box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c)) x :: ereal)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   608
      by (auto split: split_indicator)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   609
    moreover
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   610
    { have "(- c) ^ card ?B * (\<Prod>x\<in>?B. l \<bullet> x - u \<bullet> x) = 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   611
         (-1 * c) ^ card ?B * (\<Prod>x\<in>?B. -1 * (u \<bullet> x - l \<bullet> x))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   612
         by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   613
      also have "\<dots> = (-1 * -1)^card ?B * c ^ card ?B * (\<Prod>x\<in>?B. u \<bullet> x - l \<bullet> x)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   614
        unfolding setprod.distrib power_mult_distrib by (simp add: setprod_constant)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   615
      finally have "(- c) ^ card ?B * (\<Prod>x\<in>?B. l \<bullet> x - u \<bullet> x) = c ^ card ?B * (\<Prod>b\<in>?B. u \<bullet> b - l \<bullet> b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   616
        by simp }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   617
    ultimately show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   618
      using `c < 0` le
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   619
      by (auto simp: field_simps emeasure_density nn_integral_distr nn_integral_cmult
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   620
                     emeasure_lborel_box_eq inner_simps setprod_dividef setprod_constant
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   621
                     borel_measurable_indicator' emeasure_distr)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   622
  qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   623
qed simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   624
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   625
lemma lborel_real_affine:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   626
  "c \<noteq> 0 \<Longrightarrow> lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. \<bar>c\<bar>)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   627
  using lborel_affine[of c t] by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   628
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   629
lemma AE_borel_affine: 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   630
  fixes P :: "real \<Rightarrow> bool"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   631
  shows "c \<noteq> 0 \<Longrightarrow> Measurable.pred borel P \<Longrightarrow> AE x in lborel. P x \<Longrightarrow> AE x in lborel. P (t + c * x)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   632
  by (subst lborel_real_affine[where t="- t / c" and c="1 / c"])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   633
     (simp_all add: AE_density AE_distr_iff field_simps)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   634
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   635
lemma nn_integral_real_affine:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   636
  fixes c :: real assumes [measurable]: "f \<in> borel_measurable borel" and c: "c \<noteq> 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   637
  shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = \<bar>c\<bar> * (\<integral>\<^sup>+x. f (t + c * x) \<partial>lborel)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   638
  by (subst lborel_real_affine[OF c, of t])
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   639
     (simp add: nn_integral_density nn_integral_distr nn_integral_cmult)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   640
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   641
lemma lborel_integrable_real_affine:
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   642
  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   643
  assumes f: "integrable lborel f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   644
  shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   645
  using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   646
  by (subst (asm) nn_integral_real_affine[where c=c and t=t]) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   647
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   648
lemma lborel_integrable_real_affine_iff:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   649
  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   650
  shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x)) \<longleftrightarrow> integrable lborel f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   651
  using
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   652
    lborel_integrable_real_affine[of f c t]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   653
    lborel_integrable_real_affine[of "\<lambda>x. f (t + c * x)" "1/c" "-t/c"]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   654
  by (auto simp add: field_simps)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   655
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   656
lemma lborel_integral_real_affine:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   657
  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}" and c :: real
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57138
diff changeset
   658
  assumes c: "c \<noteq> 0" shows "(\<integral>x. f x \<partial> lborel) = \<bar>c\<bar> *\<^sub>R (\<integral>x. f (t + c * x) \<partial>lborel)"
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57138
diff changeset
   659
proof cases
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57138
diff changeset
   660
  assume f[measurable]: "integrable lborel f" then show ?thesis
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57138
diff changeset
   661
    using c f f[THEN borel_measurable_integrable] f[THEN lborel_integrable_real_affine, of c t]
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   662
    by (subst lborel_real_affine[OF c, of t])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   663
       (simp add: integral_density integral_distr)
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57138
diff changeset
   664
next
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57138
diff changeset
   665
  assume "\<not> integrable lborel f" with c show ?thesis
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57138
diff changeset
   666
    by (simp add: lborel_integrable_real_affine_iff not_integrable_integral_eq)
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57138
diff changeset
   667
qed
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   668
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   669
lemma divideR_right: 
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   670
  fixes x y :: "'a::real_normed_vector"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   671
  shows "r \<noteq> 0 \<Longrightarrow> y = x /\<^sub>R r \<longleftrightarrow> r *\<^sub>R y = x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   672
  using scaleR_cancel_left[of r y "x /\<^sub>R r"] by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   673
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   674
lemma lborel_has_bochner_integral_real_affine_iff:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   675
  fixes x :: "'a :: {banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   676
  shows "c \<noteq> 0 \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   677
    has_bochner_integral lborel f x \<longleftrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   678
    has_bochner_integral lborel (\<lambda>x. f (t + c * x)) (x /\<^sub>R \<bar>c\<bar>)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   679
  unfolding has_bochner_integral_iff lborel_integrable_real_affine_iff
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   680
  by (simp_all add: lborel_integral_real_affine[symmetric] divideR_right cong: conj_cong)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   681
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   682
lemma lborel_distr_uminus: "distr lborel borel uminus = (lborel :: real measure)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   683
  by (subst lborel_real_affine[of "-1" 0]) 
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   684
     (auto simp: density_1 one_ereal_def[symmetric])
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   685
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   686
lemma lborel_distr_mult: 
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   687
  assumes "(c::real) \<noteq> 0"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   688
  shows "distr lborel borel (op * c) = density lborel (\<lambda>_. inverse \<bar>c\<bar>)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   689
proof-
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   690
  have "distr lborel borel (op * c) = distr lborel lborel (op * c)" by (simp cong: distr_cong)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   691
  also from assms have "... = density lborel (\<lambda>_. inverse \<bar>c\<bar>)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   692
    by (subst lborel_real_affine[of "inverse c" 0]) (auto simp: o_def distr_density_distr)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   693
  finally show ?thesis .
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   694
qed
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   695
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   696
lemma lborel_distr_mult': 
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   697
  assumes "(c::real) \<noteq> 0"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   698
  shows "lborel = density (distr lborel borel (op * c)) (\<lambda>_. abs c)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   699
proof-
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   700
  have "lborel = density lborel (\<lambda>_. 1)" by (rule density_1[symmetric])
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   701
  also from assms have "(\<lambda>_. 1 :: ereal) = (\<lambda>_. inverse (abs c) * abs c)" by (intro ext) simp
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   702
  also have "density lborel ... = density (density lborel (\<lambda>_. inverse (abs c))) (\<lambda>_. abs c)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   703
    by (subst density_density_eq) auto
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   704
  also from assms have "density lborel (\<lambda>_. inverse (abs c)) = distr lborel borel (op * c)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   705
    by (rule lborel_distr_mult[symmetric])
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   706
  finally show ?thesis .
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   707
qed
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   708
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   709
lemma lborel_distr_plus: "distr lborel borel (op + c) = (lborel :: real measure)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   710
  by (subst lborel_real_affine[of 1 c]) (auto simp: density_1 one_ereal_def[symmetric])
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   711
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   712
interpretation lborel!: sigma_finite_measure lborel
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   713
  by (rule sigma_finite_lborel)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   714
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   715
interpretation lborel_pair: pair_sigma_finite lborel lborel ..
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   716
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   717
lemma lborel_prod:
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   718
  "lborel \<Otimes>\<^sub>M lborel = (lborel :: ('a::euclidean_space \<times> 'b::euclidean_space) measure)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   719
proof (rule lborel_eqI[symmetric], clarify)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   720
  fix la ua :: 'a and lb ub :: 'b
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   721
  assume lu: "\<And>a b. (a, b) \<in> Basis \<Longrightarrow> (la, lb) \<bullet> (a, b) \<le> (ua, ub) \<bullet> (a, b)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   722
  have [simp]:
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   723
    "\<And>b. b \<in> Basis \<Longrightarrow> la \<bullet> b \<le> ua \<bullet> b"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   724
    "\<And>b. b \<in> Basis \<Longrightarrow> lb \<bullet> b \<le> ub \<bullet> b"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   725
    "inj_on (\<lambda>u. (u, 0)) Basis" "inj_on (\<lambda>u. (0, u)) Basis"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   726
    "(\<lambda>u. (u, 0)) ` Basis \<inter> (\<lambda>u. (0, u)) ` Basis = {}"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   727
    "box (la, lb) (ua, ub) = box la ua \<times> box lb ub"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   728
    using lu[of _ 0] lu[of 0] by (auto intro!: inj_onI simp add: Basis_prod_def ball_Un box_def)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   729
  show "emeasure (lborel \<Otimes>\<^sub>M lborel) (box (la, lb) (ua, ub)) =
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   730
      ereal (setprod (op \<bullet> ((ua, ub) - (la, lb))) Basis)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   731
    by (simp add: lborel.emeasure_pair_measure_Times Basis_prod_def setprod.union_disjoint
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   732
                  setprod.reindex)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   733
qed (simp add: borel_prod[symmetric])
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   734
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   735
(* FIXME: conversion in measurable prover *)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   736
lemma lborelD_Collect[measurable (raw)]: "{x\<in>space borel. P x} \<in> sets borel \<Longrightarrow> {x\<in>space lborel. P x} \<in> sets lborel" by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   737
lemma lborelD[measurable (raw)]: "A \<in> sets borel \<Longrightarrow> A \<in> sets lborel" by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   738
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   739
subsection {* Equivalence Lebesgue integral on @{const lborel} and HK-integral *}
41706
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
   740
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   741
lemma has_integral_measure_lborel:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   742
  fixes A :: "'a::euclidean_space set"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   743
  assumes A[measurable]: "A \<in> sets borel" and finite: "emeasure lborel A < \<infinity>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   744
  shows "((\<lambda>x. 1) has_integral measure lborel A) A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   745
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   746
  { fix l u :: 'a
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   747
    have "((\<lambda>x. 1) has_integral measure lborel (box l u)) (box l u)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   748
    proof cases
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   749
      assume "\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   750
      then show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   751
        apply simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   752
        apply (subst has_integral_restrict[symmetric, OF box_subset_cbox])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   753
        apply (subst has_integral_spike_interior_eq[where g="\<lambda>_. 1"])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   754
        using has_integral_const[of "1::real" l u]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   755
        apply (simp_all add: inner_diff_left[symmetric] content_cbox_cases)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   756
        done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   757
    next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   758
      assume "\<not> (\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   759
      then have "box l u = {}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   760
        unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   761
      then show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   762
        by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   763
    qed }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   764
  note has_integral_box = this
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   765
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   766
  { fix a b :: 'a let ?M = "\<lambda>A. measure lborel (A \<inter> box a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   767
    have "Int_stable  (range (\<lambda>(a, b). box a b))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   768
      by (auto simp: Int_stable_def box_Int_box)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   769
    moreover have "(range (\<lambda>(a, b). box a b)) \<subseteq> Pow UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   770
      by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   771
    moreover have "A \<in> sigma_sets UNIV (range (\<lambda>(a, b). box a b))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   772
       using A unfolding borel_eq_box by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   773
    ultimately have "((\<lambda>x. 1) has_integral ?M A) (A \<inter> box a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   774
    proof (induction rule: sigma_sets_induct_disjoint)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   775
      case (basic A) then show ?case
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   776
        by (auto simp: box_Int_box has_integral_box)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   777
    next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   778
      case empty then show ?case
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   779
        by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   780
    next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   781
      case (compl A)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   782
      then have [measurable]: "A \<in> sets borel"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   783
        by (simp add: borel_eq_box)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   784
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   785
      have "((\<lambda>x. 1) has_integral ?M (box a b)) (box a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   786
        by (simp add: has_integral_box)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   787
      moreover have "((\<lambda>x. if x \<in> A \<inter> box a b then 1 else 0) has_integral ?M A) (box a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   788
        by (subst has_integral_restrict) (auto intro: compl)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   789
      ultimately have "((\<lambda>x. 1 - (if x \<in> A \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   790
        by (rule has_integral_sub)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   791
      then have "((\<lambda>x. (if x \<in> (UNIV - A) \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   792
        by (rule has_integral_eq_eq[THEN iffD1, rotated 1]) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   793
      then have "((\<lambda>x. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) \<inter> box a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   794
        by (subst (asm) has_integral_restrict) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   795
      also have "?M (box a b) - ?M A = ?M (UNIV - A)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   796
        by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   797
      finally show ?case .
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   798
    next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   799
      case (union F)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   800
      then have [measurable]: "\<And>i. F i \<in> sets borel"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   801
        by (simp add: borel_eq_box subset_eq)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   802
      have "((\<lambda>x. if x \<in> UNION UNIV F \<inter> box a b then 1 else 0) has_integral ?M (\<Union>i. F i)) (box a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   803
      proof (rule has_integral_monotone_convergence_increasing)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   804
        let ?f = "\<lambda>k x. \<Sum>i<k. if x \<in> F i \<inter> box a b then 1 else 0 :: real"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   805
        show "\<And>k. (?f k has_integral (\<Sum>i<k. ?M (F i))) (box a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   806
          using union.IH by (auto intro!: has_integral_setsum simp del: Int_iff)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   807
        show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   808
          by (intro setsum_mono2) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   809
        from union(1) have *: "\<And>x i j. x \<in> F i \<Longrightarrow> x \<in> F j \<longleftrightarrow> j = i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   810
          by (auto simp add: disjoint_family_on_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   811
        show "\<And>x. (\<lambda>k. ?f k x) ----> (if x \<in> UNION UNIV F \<inter> box a b then 1 else 0)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   812
          apply (auto simp: * setsum.If_cases Iio_Int_singleton)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   813
          apply (rule_tac k="Suc xa" in LIMSEQ_offset)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   814
          apply (simp add: tendsto_const)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   815
          done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   816
        have *: "emeasure lborel ((\<Union>x. F x) \<inter> box a b) \<le> emeasure lborel (box a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   817
          by (intro emeasure_mono) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   818
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   819
        with union(1) show "(\<lambda>k. \<Sum>i<k. ?M (F i)) ----> ?M (\<Union>i. F i)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   820
          unfolding sums_def[symmetric] UN_extend_simps
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   821
          by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   822
      qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   823
      then show ?case
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   824
        by (subst (asm) has_integral_restrict) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   825
    qed }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   826
  note * = this
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   827
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   828
  show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   829
  proof (rule has_integral_monotone_convergence_increasing)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   830
    let ?B = "\<lambda>n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One) :: 'a set"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   831
    let ?f = "\<lambda>n::nat. \<lambda>x. if x \<in> A \<inter> ?B n then 1 else 0 :: real"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   832
    let ?M = "\<lambda>n. measure lborel (A \<inter> ?B n)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   833
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   834
    show "\<And>n::nat. (?f n has_integral ?M n) A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   835
      using * by (subst has_integral_restrict) simp_all
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   836
    show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   837
      by (auto simp: box_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   838
    { fix x assume "x \<in> A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   839
      moreover have "(\<lambda>k. indicator (A \<inter> ?B k) x :: real) ----> indicator (\<Union>k::nat. A \<inter> ?B k) x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   840
        by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   841
      ultimately show "(\<lambda>k. if x \<in> A \<inter> ?B k then 1 else 0::real) ----> 1"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   842
        by (simp add: indicator_def UN_box_eq_UNIV) }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   843
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   844
    have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) ----> emeasure lborel (\<Union>n::nat. A \<inter> ?B n)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   845
      by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   846
    also have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) = (\<lambda>n. measure lborel (A \<inter> ?B n))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   847
    proof (intro ext emeasure_eq_ereal_measure)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   848
      fix n have "emeasure lborel (A \<inter> ?B n) \<le> emeasure lborel (?B n)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   849
        by (intro emeasure_mono) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   850
      then show "emeasure lborel (A \<inter> ?B n) \<noteq> \<infinity>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   851
        by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   852
    qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   853
    finally show "(\<lambda>n. measure lborel (A \<inter> ?B n)) ----> measure lborel A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   854
      using emeasure_eq_ereal_measure[of lborel A] finite
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   855
      by (simp add: UN_box_eq_UNIV)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   856
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   857
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   858
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   859
lemma nn_integral_has_integral:
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   860
  fixes f::"'a::euclidean_space \<Rightarrow> real"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   861
  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) = ereal r"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   862
  shows "(f has_integral r) UNIV"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   863
using f proof (induct arbitrary: r rule: borel_measurable_induct_real)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   864
  case (set A)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   865
  moreover then have "((\<lambda>x. 1) has_integral measure lborel A) A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   866
    by (intro has_integral_measure_lborel) (auto simp: ereal_indicator)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   867
  ultimately show ?case
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   868
    by (simp add: ereal_indicator measure_def) (simp add: indicator_def)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   869
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   870
  case (mult g c)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   871
  then have "ereal c * (\<integral>\<^sup>+ x. g x \<partial>lborel) = ereal r"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   872
    by (subst nn_integral_cmult[symmetric]) auto
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   873
  then obtain r' where "(c = 0 \<and> r = 0) \<or> ((\<integral>\<^sup>+ x. ereal (g x) \<partial>lborel) = ereal r' \<and> r = c * r')"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   874
    by (cases "\<integral>\<^sup>+ x. ereal (g x) \<partial>lborel") (auto split: split_if_asm)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   875
  with mult show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   876
    by (auto intro!: has_integral_cmult_real)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   877
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   878
  case (add g h)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   879
  moreover
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   880
  then have "(\<integral>\<^sup>+ x. h x + g x \<partial>lborel) = (\<integral>\<^sup>+ x. h x \<partial>lborel) + (\<integral>\<^sup>+ x. g x \<partial>lborel)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   881
    unfolding plus_ereal.simps[symmetric] by (subst nn_integral_add) auto
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   882
  with add obtain a b where "(\<integral>\<^sup>+ x. h x \<partial>lborel) = ereal a" "(\<integral>\<^sup>+ x. g x \<partial>lborel) = ereal b" "r = a + b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   883
    by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ereal2_cases) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   884
  ultimately show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   885
    by (auto intro!: has_integral_add)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   886
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   887
  case (seq U)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   888
  note seq(1)[measurable] and f[measurable]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   889
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   890
  { fix i x 
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   891
    have "U i x \<le> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   892
      using seq(5)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   893
      apply (rule LIMSEQ_le_const)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   894
      using seq(4)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   895
      apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   896
      done }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   897
  note U_le_f = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   898
  
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   899
  { fix i
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   900
    have "(\<integral>\<^sup>+x. ereal (U i x) \<partial>lborel) \<le> (\<integral>\<^sup>+x. ereal (f x) \<partial>lborel)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   901
      using U_le_f by (intro nn_integral_mono) simp
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   902
    then obtain p where "(\<integral>\<^sup>+x. U i x \<partial>lborel) = ereal p" "p \<le> r"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   903
      using seq(6) by (cases "\<integral>\<^sup>+x. U i x \<partial>lborel") auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   904
    moreover then have "0 \<le> p"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   905
      by (metis ereal_less_eq(5) nn_integral_nonneg)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   906
    moreover note seq
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   907
    ultimately have "\<exists>p. (\<integral>\<^sup>+x. U i x \<partial>lborel) = ereal p \<and> 0 \<le> p \<and> p \<le> r \<and> (U i has_integral p) UNIV"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   908
      by auto }
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   909
  then obtain p where p: "\<And>i. (\<integral>\<^sup>+x. ereal (U i x) \<partial>lborel) = ereal (p i)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   910
    and bnd: "\<And>i. p i \<le> r" "\<And>i. 0 \<le> p i"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   911
    and U_int: "\<And>i.(U i has_integral (p i)) UNIV" by metis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   912
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   913
  have int_eq: "\<And>i. integral UNIV (U i) = p i" using U_int by (rule integral_unique)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   914
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   915
  have *: "f integrable_on UNIV \<and> (\<lambda>k. integral UNIV (U k)) ----> integral UNIV f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   916
  proof (rule monotone_convergence_increasing)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   917
    show "\<forall>k. U k integrable_on UNIV" using U_int by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   918
    show "\<forall>k. \<forall>x\<in>UNIV. U k x \<le> U (Suc k) x" using `incseq U` by (auto simp: incseq_def le_fun_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   919
    then show "bounded {integral UNIV (U k) |k. True}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   920
      using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   921
    show "\<forall>x\<in>UNIV. (\<lambda>k. U k x) ----> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   922
      using seq by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   923
  qed
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   924
  moreover have "(\<lambda>i. (\<integral>\<^sup>+x. U i x \<partial>lborel)) ----> (\<integral>\<^sup>+x. f x \<partial>lborel)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   925
    using seq U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   926
  ultimately have "integral UNIV f = r"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   927
    by (auto simp add: int_eq p seq intro: LIMSEQ_unique)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   928
  with * show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   929
    by (simp add: has_integral_integral)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   930
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   931
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   932
lemma nn_integral_lborel_eq_integral:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   933
  fixes f::"'a::euclidean_space \<Rightarrow> real"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   934
  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   935
  shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = integral UNIV f"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   936
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   937
  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ereal r"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   938
    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel") auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   939
  then show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   940
    using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   941
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   942
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   943
lemma nn_integral_integrable_on:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   944
  fixes f::"'a::euclidean_space \<Rightarrow> real"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   945
  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   946
  shows "f integrable_on UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   947
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   948
  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ereal r"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   949
    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel") auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   950
  then show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   951
    by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   952
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   953
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   954
lemma nn_integral_has_integral_lborel: 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   955
  fixes f :: "'a::euclidean_space \<Rightarrow> real"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   956
  assumes f_borel: "f \<in> borel_measurable borel" and nonneg: "\<And>x. 0 \<le> f x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   957
  assumes I: "(f has_integral I) UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   958
  shows "integral\<^sup>N lborel f = I"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   959
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   960
  from f_borel have "(\<lambda>x. ereal (f x)) \<in> borel_measurable lborel" by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   961
  from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   962
  let ?B = "\<lambda>i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   963
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   964
  note F(1)[THEN borel_measurable_simple_function, measurable]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   965
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   966
  { fix i x have "real (F i x) \<le> f x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   967
      using F(3,5) F(4)[of x, symmetric] nonneg
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   968
      unfolding real_le_ereal_iff
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   969
      by (auto simp: image_iff eq_commute[of \<infinity>] max_def intro: SUP_upper split: split_if_asm) }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   970
  note F_le_f = this
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   971
  let ?F = "\<lambda>i x. F i x * indicator (?B i) x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   972
  have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lborel) = (SUP i. integral\<^sup>N lborel (\<lambda>x. ?F i x))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   973
  proof (subst nn_integral_monotone_convergence_SUP[symmetric])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   974
    { fix x
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   975
      obtain j where j: "x \<in> ?B j"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   976
        using UN_box_eq_UNIV by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   977
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   978
      have "ereal (f x) = (SUP i. F i x)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   979
        using F(4)[of x] nonneg[of x] by (simp add: max_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   980
      also have "\<dots> = (SUP i. ?F i x)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   981
      proof (rule SUP_eq)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   982
        fix i show "\<exists>j\<in>UNIV. F i x \<le> ?F j x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   983
          using j F(2)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   984
          by (intro bexI[of _ "max i j"])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   985
             (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   986
      qed (auto intro!: F split: split_indicator)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   987
      finally have "ereal (f x) =  (SUP i. ?F i x)" . }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   988
    then show "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lborel) = (\<integral>\<^sup>+ x. (SUP i. ?F i x) \<partial>lborel)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   989
      by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   990
  qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   991
  also have "\<dots> \<le> ereal I"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   992
  proof (rule SUP_least)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   993
    fix i :: nat
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   994
    have finite_F: "(\<integral>\<^sup>+ x. ereal (real (F i x) * indicator (?B i) x) \<partial>lborel) < \<infinity>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   995
    proof (rule nn_integral_bound_simple_function)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   996
      have "emeasure lborel {x \<in> space lborel. ereal (real (F i x) * indicator (?B i) x) \<noteq> 0} \<le>
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   997
        emeasure lborel (?B i)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   998
        by (intro emeasure_mono)  (auto split: split_indicator)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   999
      then show "emeasure lborel {x \<in> space lborel. ereal (real (F i x) * indicator (?B i) x) \<noteq> 0} < \<infinity>"  
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1000
        by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1001
    qed (auto split: split_indicator
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1002
              intro!: real_of_ereal_pos F simple_function_compose1[where g="real"] simple_function_ereal)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1003
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1004
    have int_F: "(\<lambda>x. real (F i x) * indicator (?B i) x) integrable_on UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1005
      using F(5) finite_F
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1006
      by (intro nn_integral_integrable_on) (auto split: split_indicator intro: real_of_ereal_pos)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1007
    
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1008
    have "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) = 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1009
      (\<integral>\<^sup>+ x. ereal (real (F i x) * indicator (?B i) x) \<partial>lborel)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1010
      using F(3,5)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1011
      by (intro nn_integral_cong) (auto simp: image_iff ereal_real eq_commute split: split_indicator)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1012
    also have "\<dots> = ereal (integral UNIV (\<lambda>x. real (F i x) * indicator (?B i) x))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1013
      using F
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1014
      by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1015
         (auto split: split_indicator intro: real_of_ereal_pos)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1016
    also have "\<dots> \<le> ereal I"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1017
      by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1018
          split: split_indicator )
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1019
    finally show "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) \<le> ereal I" .
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1020
  qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1021
  finally have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lborel) < \<infinity>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1022
    by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1023
  from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1024
    by (simp add: integral_unique)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1025
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1026
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1027
lemma has_integral_iff_emeasure_lborel:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1028
  fixes A :: "'a::euclidean_space set"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1029
  assumes A[measurable]: "A \<in> sets borel"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1030
  shows "((\<lambda>x. 1) has_integral r) A \<longleftrightarrow> emeasure lborel A = ereal r"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1031
proof cases
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1032
  assume emeasure_A: "emeasure lborel A = \<infinity>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1033
  have "\<not> (\<lambda>x. 1::real) integrable_on A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1034
  proof
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1035
    assume int: "(\<lambda>x. 1::real) integrable_on A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1036
    then have "(indicator A::'a \<Rightarrow> real) integrable_on UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1037
      unfolding indicator_def[abs_def] integrable_restrict_univ .
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1038
    then obtain r where "((indicator A::'a\<Rightarrow>real) has_integral r) UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1039
      by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1040
    from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1041
      by (simp add: ereal_indicator)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1042
  qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1043
  with emeasure_A show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1044
    by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1045
next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1046
  assume "emeasure lborel A \<noteq> \<infinity>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1047
  moreover then have "((\<lambda>x. 1) has_integral measure lborel A) A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1048
    by (simp add: has_integral_measure_lborel)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1049
  ultimately show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1050
    by (auto simp: emeasure_eq_ereal_measure has_integral_unique)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1051
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1052
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1053
lemma has_integral_integral_real:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1054
  fixes f::"'a::euclidean_space \<Rightarrow> real"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1055
  assumes f: "integrable lborel f"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1056
  shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1057
using f proof induct
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1058
  case (base A c) then show ?case
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1059
    by (auto intro!: has_integral_mult_left simp: )
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1060
       (simp add: emeasure_eq_ereal_measure indicator_def has_integral_measure_lborel)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1061
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1062
  case (add f g) then show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1063
    by (auto intro!: has_integral_add)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1064
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1065
  case (lim f s)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1066
  show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1067
  proof (rule has_integral_dominated_convergence)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1068
    show "\<And>i. (s i has_integral integral\<^sup>L lborel (s i)) UNIV" by fact
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1069
    show "(\<lambda>x. norm (2 * f x)) integrable_on UNIV"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1070
      using `integrable lborel f`
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1071
      by (intro nn_integral_integrable_on)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1072
         (auto simp: integrable_iff_bounded abs_mult times_ereal.simps(1)[symmetric] nn_integral_cmult
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1073
               simp del: times_ereal.simps)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1074
    show "\<And>k. \<forall>x\<in>UNIV. norm (s k x) \<le> norm (2 * f x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1075
      using lim by (auto simp add: abs_mult)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1076
    show "\<forall>x\<in>UNIV. (\<lambda>k. s k x) ----> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1077
      using lim by auto
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1078
    show "(\<lambda>k. integral\<^sup>L lborel (s k)) ----> integral\<^sup>L lborel f"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1079
      using lim lim(1)[THEN borel_measurable_integrable]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1080
      by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1081
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
  1082
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
  1083
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1084
context
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1085
  fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1086
begin
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
  1087
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1088
lemma has_integral_integral_lborel:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1089
  assumes f: "integrable lborel f"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
  1090
  shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
  1091
proof -
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1092
  have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1093
    using f by (intro has_integral_setsum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1094
  also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1095
    by (simp add: fun_eq_iff euclidean_representation)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1096
  also have "(\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lborel f"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1097
    using f by (subst (2) eq_f[symmetric]) simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1098
  finally show ?thesis .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1099
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1100
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1101
lemma integrable_on_lborel: "integrable lborel f \<Longrightarrow> f integrable_on UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1102
  using has_integral_integral_lborel by (auto intro: has_integral_integrable)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1103
  
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1104
lemma integral_lborel: "integrable lborel f \<Longrightarrow> integral UNIV f = (\<integral>x. f x \<partial>lborel)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1105
  using has_integral_integral_lborel by auto
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1106
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1107
end
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1108
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1109
subsection {* Fundamental Theorem of Calculus for the Lebesgue integral *}
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1110
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1111
lemma emeasure_bounded_finite:
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1112
  assumes "bounded A" shows "emeasure lborel A < \<infinity>"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1113
proof -
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1114
  from bounded_subset_cbox[OF `bounded A`] obtain a b where "A \<subseteq> cbox a b"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1115
    by auto
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1116
  then have "emeasure lborel A \<le> emeasure lborel (cbox a b)"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1117
    by (intro emeasure_mono) auto
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1118
  then show ?thesis
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1119
    by (auto simp: emeasure_lborel_cbox_eq)
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1120
qed
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1121
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1122
lemma emeasure_compact_finite: "compact A \<Longrightarrow> emeasure lborel A < \<infinity>"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1123
  using emeasure_bounded_finite[of A] by (auto intro: compact_imp_bounded)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1124
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1125
lemma borel_integrable_compact:
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1126
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}"
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1127
  assumes "compact S" "continuous_on S f"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1128
  shows "integrable lborel (\<lambda>x. indicator S x *\<^sub>R f x)"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1129
proof cases
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1130
  assume "S \<noteq> {}"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1131
  have "continuous_on S (\<lambda>x. norm (f x))"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1132
    using assms by (intro continuous_intros)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1133
  from continuous_attains_sup[OF `compact S` `S \<noteq> {}` this]
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1134
  obtain M where M: "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> M"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1135
    by auto
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1136
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1137
  show ?thesis
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1138
  proof (rule integrable_bound)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1139
    show "integrable lborel (\<lambda>x. indicator S x * M)"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1140
      using assms by (auto intro!: emeasure_compact_finite borel_compact integrable_mult_left)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1141
    show "(\<lambda>x. indicator S x *\<^sub>R f x) \<in> borel_measurable lborel"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1142
      using assms by (auto intro!: borel_measurable_continuous_on_indicator borel_compact)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1143
    show "AE x in lborel. norm (indicator S x *\<^sub>R f x) \<le> norm (indicator S x * M)"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1144
      by (auto split: split_indicator simp: abs_real_def dest!: M)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1145
  qed
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1146
qed simp
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1147
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1148
lemma borel_integrable_atLeastAtMost:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1149
  fixes f :: "real \<Rightarrow> real"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1150
  assumes f: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1151
  shows "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is "integrable _ ?f")
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1152
proof -
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1153
  have "integrable lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x)"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1154
  proof (rule borel_integrable_compact)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1155
    from f show "continuous_on {a..b} f"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1156
      by (auto intro: continuous_at_imp_continuous_on)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1157
  qed simp
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1158
  then show ?thesis
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57447
diff changeset
  1159
    by (auto simp: mult.commute)
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1160
qed
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1161
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1162
text {*
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1163
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1164
For the positive integral we replace continuity with Borel-measurability. 
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1165
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1166
*}
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1167
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1168
lemma
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1169
  fixes f :: "real \<Rightarrow> real"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1170
  assumes [measurable]: "f \<in> borel_measurable borel"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1171
  assumes f: "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" and "a \<le> b"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1172
  shows nn_integral_FTC_Icc: "(\<integral>\<^sup>+x. ereal (f x) * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?nn)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1173
    and has_bochner_integral_FTC_Icc_nonneg:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1174
      "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1175
    and integral_FTC_Icc_nonneg: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1176
    and integrable_FTC_Icc_nonneg: "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is ?int)
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1177
proof -
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1178
  have *: "(\<lambda>x. f x * indicator {a..b} x) \<in> borel_measurable borel" "\<And>x. 0 \<le> f x * indicator {a..b} x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1179
    using f(2) by (auto split: split_indicator)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1180
    
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1181
  have "(f has_integral F b - F a) {a..b}"
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1182
    by (intro fundamental_theorem_of_calculus)
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1183
       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric]
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1184
             intro: has_field_derivative_subset[OF f(1)] `a \<le> b`)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1185
  then have i: "((\<lambda>x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1186
    unfolding indicator_def if_distrib[where f="\<lambda>x. a * x" for a]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1187
    by (simp cong del: if_cong del: atLeastAtMost_iff)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1188
  then have nn: "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1189
    by (rule nn_integral_has_integral_lborel[OF *])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1190
  then show ?has
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1191
    by (rule has_bochner_integral_nn_integral[rotated 2]) (simp_all add: *)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1192
  then show ?eq ?int
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1193
    unfolding has_bochner_integral_iff by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1194
  from nn show ?nn
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1195
    by (simp add: ereal_mult_indicator)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1196
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1197
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1198
lemma
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1199
  fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1200
  assumes "a \<le> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1201
  assumes "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1202
  assumes cont: "continuous_on {a .. b} f"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1203
  shows has_bochner_integral_FTC_Icc:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1204
      "has_bochner_integral lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) (F b - F a)" (is ?has)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1205
    and integral_FTC_Icc: "(\<integral>x. indicator {a .. b} x *\<^sub>R f x \<partial>lborel) = F b - F a" (is ?eq)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1206
proof -
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1207
  let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1208
  have int: "integrable lborel ?f"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1209
    using borel_integrable_compact[OF _ cont] by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1210
  have "(f has_integral F b - F a) {a..b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1211
    using assms(1,2) by (intro fundamental_theorem_of_calculus) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1212
  moreover 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1213
  have "(f has_integral integral\<^sup>L lborel ?f) {a..b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1214
    using has_integral_integral_lborel[OF int]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1215
    unfolding indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1216
    by (simp cong del: if_cong del: atLeastAtMost_iff)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1217
  ultimately show ?eq
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1218
    by (auto dest: has_integral_unique)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1219
  then show ?has
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1220
    using int by (auto simp: has_bochner_integral_iff)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1221
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1222
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1223
lemma
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1224
  fixes f :: "real \<Rightarrow> real"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1225
  assumes "a \<le> b"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1226
  assumes deriv: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> DERIV F x :> f x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1227
  assumes cont: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1228
  shows has_bochner_integral_FTC_Icc_real:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1229
      "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1230
    and integral_FTC_Icc_real: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1231
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1232
  have 1: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1233
    unfolding has_field_derivative_iff_has_vector_derivative[symmetric]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1234
    using deriv by (auto intro: DERIV_subset)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1235
  have 2: "continuous_on {a .. b} f"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1236
    using cont by (intro continuous_at_imp_continuous_on) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1237
  show ?has ?eq
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1238
    using has_bochner_integral_FTC_Icc[OF `a \<le> b` 1 2] integral_FTC_Icc[OF `a \<le> b` 1 2]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57447
diff changeset
  1239
    by (auto simp: mult.commute)
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1240
qed
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1241
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1242
lemma nn_integral_FTC_atLeast:
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1243
  fixes f :: "real \<Rightarrow> real"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1244
  assumes f_borel: "f \<in> borel_measurable borel"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1245
  assumes f: "\<And>x. a \<le> x \<Longrightarrow> DERIV F x :> f x" 
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1246
  assumes nonneg: "\<And>x. a \<le> x \<Longrightarrow> 0 \<le> f x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1247
  assumes lim: "(F ---> T) at_top"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
  1248
  shows "(\<integral>\<^sup>+x. ereal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1249
proof -
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1250
  let ?f = "\<lambda>(i::nat) (x::real). ereal (f x) * indicator {a..a + real i} x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1251
  let ?fR = "\<lambda>x. ereal (f x) * indicator {a ..} x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1252
  have "\<And>x. (SUP i::nat. ?f i x) = ?fR x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1253
  proof (rule SUP_Lim_ereal)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1254
    show "\<And>x. incseq (\<lambda>i. ?f i x)"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1255
      using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1256
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1257
    fix x
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1258
    from reals_Archimedean2[of "x - a"] guess n ..
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1259
    then have "eventually (\<lambda>n. ?f n x = ?fR x) sequentially"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1260
      by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1261
    then show "(\<lambda>n. ?f n x) ----> ?fR x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1262
      by (rule Lim_eventually)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1263
  qed
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1264
  then have "integral\<^sup>N lborel ?fR = (\<integral>\<^sup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1265
    by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
  1266
  also have "\<dots> = (SUP i::nat. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1267
  proof (rule nn_integral_monotone_convergence_SUP)
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1268
    show "incseq ?f"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1269
      using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1270
    show "\<And>i. (?f i) \<in> borel_measurable lborel"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1271
      using f_borel by auto
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1272
  qed
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 53374
diff changeset
  1273
  also have "\<dots> = (SUP i::nat. ereal (F (a + real i) - F a))"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1274
    by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1275
  also have "\<dots> = T - F a"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1276
  proof (rule SUP_Lim_ereal)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1277
    show "incseq (\<lambda>n. ereal (F (a + real n) - F a))"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1278
    proof (simp add: incseq_def, safe)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1279
      fix m n :: nat assume "m \<le> n"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1280
      with f nonneg show "F (a + real m) \<le> F (a + real n)"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1281
        by (intro DERIV_nonneg_imp_nondecreasing[where f=F])
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1282
           (simp, metis add_increasing2 order_refl order_trans real_of_nat_ge_zero)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1283
    qed 
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1284
    have "(\<lambda>x. F (a + real x)) ----> T"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1285
      apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top])
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1286
      apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl])
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1287
      apply (rule filterlim_real_sequentially)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1288
      done
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1289
    then show "(\<lambda>n. ereal (F (a + real n) - F a)) ----> ereal (T - F a)"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1290
      unfolding lim_ereal
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1291
      by (intro tendsto_diff) auto
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1292
  qed
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1293
  finally show ?thesis .
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1294
qed
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1295
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1296
lemma integral_power:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1297
  "a \<le> b \<Longrightarrow> (\<integral>x. x^k * indicator {a..b} x \<partial>lborel) = (b^Suc k - a^Suc k) / Suc k"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1298
proof (subst integral_FTC_Icc_real)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1299
  fix x show "DERIV (\<lambda>x. x^Suc k / Suc k) x :> x^k"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1300
    by (intro derivative_eq_intros) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1301
qed (auto simp: field_simps)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1302
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1303
subsection {* Integration by parts *}
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1304
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1305
lemma integral_by_parts_integrable:
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1306
  fixes f g F G::"real \<Rightarrow> real"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1307
  assumes "a \<le> b"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1308
  assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1309
  assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1310
  assumes [intro]: "!!x. DERIV F x :> f x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1311
  assumes [intro]: "!!x. DERIV G x :> g x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1312
  shows  "integrable lborel (\<lambda>x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1313
  by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1314
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1315
lemma integral_by_parts:
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1316
  fixes f g F G::"real \<Rightarrow> real"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1317
  assumes [arith]: "a \<le> b"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1318
  assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1319
  assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1320
  assumes [intro]: "!!x. DERIV F x :> f x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1321
  assumes [intro]: "!!x. DERIV G x :> g x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1322
  shows "(\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1323
            =  F b * G b - F a * G a - \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel" 
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1324
proof-
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1325
  have 0: "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) = F b * G b - F a * G a"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1326
    by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros) 
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1327
      (auto intro!: DERIV_isCont)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1328
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1329
  have "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) =
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1330
    (\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel) + \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1331
    apply (subst integral_add[symmetric])
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1332
    apply (auto intro!: borel_integrable_atLeastAtMost continuous_intros)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1333
    by (auto intro!: DERIV_isCont integral_cong split:split_indicator)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1334
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1335
  thus ?thesis using 0 by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1336
qed
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1337
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1338
lemma integral_by_parts':
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1339
  fixes f g F G::"real \<Rightarrow> real"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1340
  assumes "a \<le> b"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1341
  assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1342
  assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1343
  assumes "!!x. DERIV F x :> f x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1344
  assumes "!!x. DERIV G x :> g x"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1345
  shows "(\<integral>x. indicator {a .. b} x *\<^sub>R (F x * g x) \<partial>lborel)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1346
            =  F b * G b - F a * G a - \<integral>x. indicator {a .. b} x *\<^sub>R (f x * G x) \<partial>lborel" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1347
  using integral_by_parts[OF assms] by (simp add: ac_simps)
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1348
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1349
lemma has_bochner_integral_even_function:
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1350
  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1351
  assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1352
  assumes even: "\<And>x. f (- x) = f x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1353
  shows "has_bochner_integral lborel f (2 *\<^sub>R x)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1354
proof -
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1355
  have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1356
    by (auto split: split_indicator)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1357
  have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1358
    by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1359
       (auto simp: indicator even f)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1360
  with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1361
    by (rule has_bochner_integral_add)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1362
  then have "has_bochner_integral lborel f (x + x)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1363
    by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1364
       (auto split: split_indicator)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1365
  then show ?thesis
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1366
    by (simp add: scaleR_2)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1367
qed
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1368
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1369
lemma has_bochner_integral_odd_function:
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1370
  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1371
  assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1372
  assumes odd: "\<And>x. f (- x) = - f x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1373
  shows "has_bochner_integral lborel f 0"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1374
proof -
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1375
  have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1376
    by (auto split: split_indicator)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1377
  have "has_bochner_integral lborel (\<lambda>x. - indicator {.. 0} x *\<^sub>R f x) x"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1378
    by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1379
       (auto simp: indicator odd f)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1380
  from has_bochner_integral_minus[OF this]
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1381
  have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) (- x)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1382
    by simp 
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1383
  with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1384
    by (rule has_bochner_integral_add)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1385
  then have "has_bochner_integral lborel f (x + - x)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1386
    by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1387
       (auto split: split_indicator)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1388
  then show ?thesis
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1389
    by simp
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1390
qed
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57166
diff changeset
  1391
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1392
end
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1393