src/HOL/Analysis/Lebesgue_Measure.thy
author paulson <lp15@cam.ac.uk>
Thu, 12 Sep 2019 14:51:45 +0100
changeset 70688 3d894e1cfc75
parent 70547 7ce95a5c4aa8
child 70694 ae37b8fbf023
permissions -rw-r--r--
new material on Analysis, plus some rearrangements
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63626
diff changeset
     1
(*  Title:      HOL/Analysis/Lebesgue_Measure.thy
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     2
    Author:     Johannes Hölzl, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     3
    Author:     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
69517
dc20f278e8f3 tuned style and headers
nipkow
parents: 69447
diff changeset
     8
section \<open>Lebesgue Measure\<close>
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     9
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    10
theory Lebesgue_Measure
69517
dc20f278e8f3 tuned style and headers
nipkow
parents: 69447
diff changeset
    11
imports
dc20f278e8f3 tuned style and headers
nipkow
parents: 69447
diff changeset
    12
  Finite_Product_Measure
dc20f278e8f3 tuned style and headers
nipkow
parents: 69447
diff changeset
    13
  Caratheodory
dc20f278e8f3 tuned style and headers
nipkow
parents: 69447
diff changeset
    14
  Complete_Measure
dc20f278e8f3 tuned style and headers
nipkow
parents: 69447
diff changeset
    15
  Summation_Tests
dc20f278e8f3 tuned style and headers
nipkow
parents: 69447
diff changeset
    16
  Regularity
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    17
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    18
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    19
lemma measure_eqI_lessThan:
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    20
  fixes M N :: "real measure"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    21
  assumes sets: "sets M = sets borel" "sets N = sets borel"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    22
  assumes fin: "\<And>x. emeasure M {x <..} < \<infinity>"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    23
  assumes "\<And>x. emeasure M {x <..} = emeasure N {x <..}"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    24
  shows "M = N"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    25
proof (rule measure_eqI_generator_eq_countable)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    26
  let ?LT = "\<lambda>a::real. {a <..}" let ?E = "range ?LT"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    27
  show "Int_stable ?E"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    28
    by (auto simp: Int_stable_def lessThan_Int_lessThan)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    29
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    30
  show "?E \<subseteq> Pow UNIV" "sets M = sigma_sets UNIV ?E" "sets N = sigma_sets UNIV ?E"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    31
    unfolding sets borel_Ioi by auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    32
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    33
  show "?LT`Rats \<subseteq> ?E" "(\<Union>i\<in>Rats. ?LT i) = UNIV" "\<And>a. a \<in> ?LT`Rats \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    34
    using fin by (auto intro: Rats_no_bot_less simp: less_top)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    35
qed (auto intro: assms countable_rat)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63968
diff changeset
    36
69447
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
    37
subsection \<open>Measures defined by monotonous functions\<close>
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
    38
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
    39
text \<open>
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
    40
  Every right-continuous and nondecreasing function gives rise to a measure on the reals:
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
    41
\<close>
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    42
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
    43
definition\<^marker>\<open>tag important\<close> interval_measure :: "(real \<Rightarrow> real) \<Rightarrow> real measure" where
69447
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
    44
  "interval_measure F =
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
    45
     extend_measure UNIV {(a, b). a \<le> b} (\<lambda>(a, b). {a<..b}) (\<lambda>(a, b). ennreal (F b - F a))"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    46
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
    47
lemma\<^marker>\<open>tag important\<close> emeasure_interval_measure_Ioc:
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    48
  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
    49
  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
    50
  assumes right_cont_F : "\<And>a. continuous (at_right a) F"
69447
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
    51
  shows "emeasure (interval_measure F) {a<..b} = F b - F a"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
    52
proof (rule extend_measure_caratheodory_pair[OF interval_measure_def \<open>a \<le> b\<close>])
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    53
  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
    54
  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
    55
    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
    56
    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
    57
    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
    58
      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
    59
      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
    60
      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
    61
        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
    62
      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
    63
    next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    64
      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
    65
      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
    66
      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
    67
        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
    68
      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
    69
    qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    70
  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
    71
next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    72
  fix l r :: "nat \<Rightarrow> real" and a b :: real
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
    73
  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})"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    74
  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
    75
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61610
diff changeset
    76
  have [intro, simp]: "\<And>a b. a \<le> b \<Longrightarrow> F a \<le> F b"
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61610
diff changeset
    77
    by (auto intro!: l_r mono_F)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    78
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    79
  { fix S :: "nat set" assume "finite S"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
    80
    moreover note \<open>a \<le> b\<close>
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    81
    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
    82
      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
    83
    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
    84
    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
    85
      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
    86
      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
    87
      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
    88
        assume "\<exists>i\<in>S. l i < r i"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
    89
        with \<open>finite S\<close> have "Min (l ` {i\<in>S. l i < r i}) \<in> l ` {i\<in>S. l i < r 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
    90
          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
    91
        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
    92
          by fastforce
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    93
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
    94
        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))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
    95
          using m psubset by (intro sum.remove) 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
    96
        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
    97
        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
    98
          show "S - {m} \<subset> S"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
    99
            using \<open>m\<in>S\<close> 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
   100
          show "r m \<le> b"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   101
            using psubset.prems(2)[OF \<open>m\<in>S\<close>] \<open>l m < r m\<close> 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
   102
        next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   103
          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
   104
          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
   105
          { assume i': "l i < r i" "l i < r m"
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63262
diff changeset
   106
            with \<open>finite S\<close> i m have "l m \<le> l 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
   107
              by auto
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63262
diff changeset
   108
            with i' have "{l i <.. r i} \<inter> {l m <.. r m} \<noteq> {}"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   109
              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
   110
            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
   111
              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
   112
          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
   113
            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
   114
          then show "{l i <.. r i} \<subseteq> {r m <.. b}"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   115
            using psubset.prems(2)[OF \<open>i\<in>S\<close>] 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
   116
        qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   117
        also have "F (r m) - F (l m) \<le> F (r m) - F a"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   118
          using psubset.prems(2)[OF \<open>m \<in> S\<close>] \<open>l m < r m\<close>
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   119
          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
   120
        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
   121
          by (auto intro: add_mono)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   122
      qed (auto simp add: \<open>a \<le> b\<close> less_le)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   123
    qed }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   124
  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
   125
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   126
  (* 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
   127
     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
   128
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   129
  { 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
   130
    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
   131
    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
   132
    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
   133
      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
   134
      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
   135
      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
   136
        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
   137
          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
   138
      next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   139
        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
   140
        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
   141
          by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   142
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   143
        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
   144
        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
   145
        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
   146
          assume "?R"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   147
          with \<open>j \<in> S\<close> psubset.prems have "{u..v} \<subseteq> (\<Union>i\<in>S-{j}. {l i<..< r 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
   148
            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
   149
            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
   150
            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
   151
            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
   152
            done
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   153
          with \<open>j \<in> S\<close> have "F v - F u \<le> (\<Sum>i\<in>S - {j}. F (r i) - F (l 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
   154
            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
   155
          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
   156
            using psubset.prems
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   157
            by (intro sum_mono2 psubset) (auto intro: less_imp_le)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   158
          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
   159
        next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   160
          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
   161
          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
   162
            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
   163
          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
   164
          let ?S2 = "{i \<in> S. r i > r j}"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   165
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   166
          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))"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   167
            using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   168
            by (intro sum_mono2) (auto intro: less_imp_le)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   169
          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
   170
            (\<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
   171
            using psubset(1) psubset.prems(1) j
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   172
            apply (subst sum.union_disjoint)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   173
            apply simp_all
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   174
            apply (subst sum.union_disjoint)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   175
            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
   176
            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
   177
            done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   178
          also (xtrans) have "(\<Sum>i\<in>?S1. F (r i) - F (l i)) \<ge> F (l j) - F u"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   179
            using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
57447
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 (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
   181
            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
   182
            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
   183
            done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   184
          also (xtrans) have "(\<Sum>i\<in>?S2. F (r i) - F (l i)) \<ge> F v - F (r j)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   185
            using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   186
            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
   187
            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
   188
            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
   189
            done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   190
          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
   191
            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
   192
        qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   193
      qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   194
    qed }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   195
  note claim2 = this
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   196
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   197
  (* now prove the inequality going the other way *)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   198
  have "ennreal (F b - F a) \<le> (\<Sum>i. ennreal (F (r i) - F (l i)))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   199
  proof (rule ennreal_le_epsilon)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   200
    fix epsilon :: real assume egt0: "epsilon > 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   201
    have "\<forall>i. \<exists>d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   202
    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
   203
      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
   204
      note right_cont_F [of "r i"]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   205
      thus "\<exists>d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   206
        apply -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   207
        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
   208
        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
   209
        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
   210
        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
   211
        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
   212
    qed
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   213
    then obtain delta where
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   214
        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
   215
        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
   216
      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
   217
    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
   218
      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
   219
      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
   220
      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
   221
      apply (drule_tac x = "epsilon / 2" in spec)
59554
4044f53326c9 inlined rules to free user-space from technical names
haftmann
parents: 59425
diff changeset
   222
      using egt0 unfolding mult.commute [of 2] by force
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   223
    then obtain a' where a'lea [arith]: "a' > a" and
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   224
      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
   225
      by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   226
    define S' where "S' = {i. l i < r i}"
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   227
    obtain S :: "nat set" where
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   228
      "S \<subseteq> S'" and finS: "finite S" and
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   229
      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
   230
    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
   231
      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
   232
        by (rule compact_Icc)
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65204
diff changeset
   233
      show "\<And>i. i \<in> S' \<Longrightarrow> open ({l i<..<r i + delta i})" 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
   234
      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
   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 "{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
   237
        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
   238
      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
   239
        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
   240
        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
   241
        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
   242
        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
   243
        done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   244
      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
   245
    qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   246
    with S'_def have Sprop2: "\<And>i. i \<in> S \<Longrightarrow> l i < r i" by auto
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   247
    from finS have "\<exists>n. \<forall>i \<in> S. i \<le> n"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   248
      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
   249
    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
   250
    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
   251
      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
   252
      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
   253
      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
   254
      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
   255
      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
   256
      by (rule deltai_gt0)
61954
1d43f86f48be more symbols;
wenzelm
parents: 61945
diff changeset
   257
    also have "... \<le> (\<Sum>i \<in> S. F(r i) - F(l i) + epsilon / 2^(i+2))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   258
      apply (rule sum_mono)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   259
      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
   260
      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
   261
      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
   262
      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
   263
      by auto
61954
1d43f86f48be more symbols;
wenzelm
parents: 61945
diff changeset
   264
    also have "... = (\<Sum>i \<in> S. F(r i) - F(l i)) +
1d43f86f48be more symbols;
wenzelm
parents: 61945
diff changeset
   265
        (epsilon / 4) * (\<Sum>i \<in> S. (1 / 2)^i)" (is "_ = ?t + _")
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   266
      by (subst sum.distrib) (simp add: field_simps sum_distrib_left)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   267
    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
   268
      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
   269
      apply (rule mult_left_mono)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   270
      apply (rule sum_mono2)
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   271
      using egt0 apply 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
   272
      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
   273
    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
   274
      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
   275
      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
   276
      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
   277
      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
   278
      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
   279
      done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   280
    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
   281
      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
   282
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   283
    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
   284
      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
   285
    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
   286
      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
   287
    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
   288
      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
   289
      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
   290
      done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   291
    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
   292
      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
   293
    also have "... \<le> (\<Sum>i\<le>n. F (r i) - F (l i)) + epsilon"
65680
378a2f11bec9 Simplification of some proofs. Also key lemmas using !! rather than ! in premises
paulson <lp15@cam.ac.uk>
parents: 65585
diff changeset
   294
      using finS Sbound Sprop by (auto intro!: add_right_mono sum_mono2)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   295
    finally have "ennreal (F b - F a) \<le> (\<Sum>i\<le>n. ennreal (F (r i) - F (l i))) + epsilon"
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68120
diff changeset
   296
      using egt0 by (simp add: sum_nonneg flip: ennreal_plus)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   297
    then show "ennreal (F b - F a) \<le> (\<Sum>i. ennreal (F (r i) - F (l i))) + (epsilon :: real)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   298
      by (rule order_trans) (auto intro!: add_mono sum_le_suminf simp del: sum_ennreal)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   299
  qed
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   300
  moreover have "(\<Sum>i. ennreal (F (r i) - F (l i))) \<le> ennreal (F b - F a)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   301
    using \<open>a \<le> b\<close> by (auto intro!: suminf_le_const ennreal_le_iff[THEN iffD2] claim1)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   302
  ultimately show "(\<Sum>n. ennreal (F (r n) - F (l n))) = ennreal (F b - F a)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   303
    by (rule antisym[rotated])
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61610
diff changeset
   304
qed (auto simp: Ioc_inj mono_F)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   305
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   306
lemma measure_interval_measure_Ioc:
70271
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   307
  assumes "a \<le> b" and "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y" and "\<And>a. continuous (at_right a) F"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   308
  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
   309
  unfolding measure_def
70271
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   310
  by (simp add: assms emeasure_interval_measure_Ioc)
57447
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_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
   313
  "(\<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
   314
    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
   315
  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
   316
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   317
lemma\<^marker>\<open>tag important\<close> sets_interval_measure [simp, measurable_cong]:
69447
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
   318
    "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
   319
  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
   320
  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
   321
  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
   322
  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
   323
  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
   324
  done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   325
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   326
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
   327
  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
   328
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   329
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
   330
  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
   331
  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   332
  assumes cont_F : "continuous_on UNIV F"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   333
  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
   334
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
   335
  { 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
   336
      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
   337
      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
   338
         (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
   339
  note * = this
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   340
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   341
  let ?F = "interval_measure F"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   342
  show "((\<lambda>a. F b - F a) \<longlongrightarrow> emeasure ?F {a..b}) (at_left 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
   343
  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
   344
    show "a - 1 < a" by simp
61969
e01015e49041 more symbols;
wenzelm
parents: 61954
diff changeset
   345
    fix X assume "\<And>n. X n < a" "incseq X" "X \<longlonglongrightarrow> a"
e01015e49041 more symbols;
wenzelm
parents: 61954
diff changeset
   346
    with \<open>a \<le> b\<close> have "(\<lambda>n. emeasure ?F {X n<..b}) \<longlonglongrightarrow> emeasure ?F (\<Inter>n. {X n <..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
   347
      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
   348
      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
   349
      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
   350
      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
   351
      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
   352
      done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   353
    also have "(\<Inter>n. {X n <..b}) = {a..b}"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   354
      using \<open>\<And>n. X n < a\<close>
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   355
      apply auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61954
diff changeset
   356
      apply (rule LIMSEQ_le_const2[OF \<open>X \<longlonglongrightarrow> a\<close>])
57447
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 (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
   358
      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
   359
      done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   360
    also have "(\<lambda>n. emeasure ?F {X n<..b}) = (\<lambda>n. F b - F (X n))"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   361
      using \<open>\<And>n. X n < a\<close> \<open>a \<le> b\<close> by (subst *) (auto intro: less_imp_le less_le_trans)
61969
e01015e49041 more symbols;
wenzelm
parents: 61954
diff changeset
   362
    finally show "(\<lambda>n. F b - F (X n)) \<longlonglongrightarrow> emeasure ?F {a..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
   363
  qed
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   364
  show "((\<lambda>a. ennreal (F b - F a)) \<longlongrightarrow> F b - F a) (at_left a)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   365
    by (rule continuous_on_tendsto_compose[where g="\<lambda>x. x" and s=UNIV])
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   366
       (auto simp: continuous_on_ennreal continuous_on_diff cont_F continuous_on_const)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   367
qed (rule trivial_limit_at_left_real)
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   368
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   369
lemma\<^marker>\<open>tag important\<close> sigma_finite_interval_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
   370
  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   371
  assumes right_cont_F : "\<And>a. continuous (at_right a) F"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   372
  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
   373
  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
   374
  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
   375
  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
   376
  done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   377
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   378
subsection \<open>Lebesgue-Borel measure\<close>
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   379
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   380
definition\<^marker>\<open>tag important\<close> lborel :: "('a :: euclidean_space) measure" where
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   381
  "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
   382
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   383
abbreviation\<^marker>\<open>tag important\<close> lebesgue :: "'a::euclidean_space measure"
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
   384
  where "lebesgue \<equiv> completion lborel"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
   385
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   386
abbreviation\<^marker>\<open>tag important\<close> lebesgue_on :: "'a set \<Rightarrow> 'a::euclidean_space measure"
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
   387
  where "lebesgue_on \<Omega> \<equiv> restrict_space (completion lborel) \<Omega>"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
   388
70380
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   389
lemma lebesgue_on_mono:
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   390
  assumes major: "AE x in lebesgue_on S. P x" and minor: "\<And>x.\<lbrakk>P x; x \<in> S\<rbrakk> \<Longrightarrow> Q x"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   391
  shows "AE x in lebesgue_on S. Q x"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   392
proof -
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   393
  have "AE a in lebesgue_on S. P a \<longrightarrow> Q a"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   394
    using minor space_restrict_space by fastforce
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   395
  then show ?thesis
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   396
    using major by auto
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   397
qed
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   398
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
   399
lemma integral_eq_zero_null_sets:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
   400
  assumes "S \<in> null_sets lebesgue"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
   401
  shows "integral\<^sup>L (lebesgue_on S) f = 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
   402
proof (rule integral_eq_zero_AE)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
   403
  show "AE x in lebesgue_on S. f x = 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
   404
    by (metis (no_types, lifting) assms AE_not_in lebesgue_on_mono null_setsD2 null_sets_restrict_space order_refl)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
   405
qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
   406
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   407
lemma
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59000
diff changeset
   408
  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
   409
    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
   410
    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
   411
    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
   412
  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
   413
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
   414
lemma space_lebesgue_on [simp]: "space (lebesgue_on S) = S"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
   415
  by (simp add: space_restrict_space)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
   416
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   417
lemma sets_lebesgue_on_refl [iff]: "S \<in> sets (lebesgue_on S)"
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   418
    by (metis inf_top.right_neutral sets.top space_borel space_completion space_lborel space_restrict_space)
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   419
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   420
lemma Compl_in_sets_lebesgue: "-A \<in> sets lebesgue \<longleftrightarrow> A  \<in> sets lebesgue"
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   421
  by (metis Compl_eq_Diff_UNIV double_compl space_borel space_completion space_lborel Sigma_Algebra.sets.compl_sets)
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   422
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   423
lemma measurable_lebesgue_cong:
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   424
  assumes "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   425
  shows "f \<in> measurable (lebesgue_on S) M \<longleftrightarrow> g \<in> measurable (lebesgue_on S) M"
69546
27dae626822b prefer naming convention from datatype package for strong congruence rules
haftmann
parents: 69517
diff changeset
   426
  by (metis (mono_tags, lifting) IntD1 assms measurable_cong_simp space_restrict_space)
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   427
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   428
lemma lebesgue_on_UNIV_eq: "lebesgue_on UNIV = lebesgue"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   429
proof -
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   430
  have "measure_of UNIV (sets lebesgue) (emeasure lebesgue) = lebesgue"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   431
    by (metis measure_of_of_measure space_borel space_completion space_lborel)
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   432
  then show ?thesis
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   433
    by (auto simp: restrict_space_def)
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   434
qed
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   435
70271
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   436
lemma integrable_lebesgue_on_UNIV_eq:
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   437
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}"
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   438
  shows "integrable (lebesgue_on UNIV) f = integrable lebesgue f"
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   439
  by (auto simp: integrable_restrict_space)
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   440
lemma integral_restrict_Int:
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   441
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   442
  assumes "S \<in> sets lebesgue" "T \<in> sets lebesgue"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   443
  shows "integral\<^sup>L (lebesgue_on T) (\<lambda>x. if x \<in> S then f x else 0) = integral\<^sup>L (lebesgue_on (S \<inter> T)) f"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   444
proof -
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   445
  have "(\<lambda>x. indicat_real T x *\<^sub>R (if x \<in> S then f x else 0)) = (\<lambda>x. indicat_real (S \<inter> T) x *\<^sub>R f x)"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   446
    by (force simp: indicator_def)
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   447
  then show ?thesis
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   448
    by (simp add: assms sets.Int Bochner_Integration.integral_restrict_space)
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   449
qed
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   450
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   451
lemma integral_restrict:
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   452
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   453
  assumes "S \<subseteq> T" "S \<in> sets lebesgue" "T \<in> sets lebesgue"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   454
  shows "integral\<^sup>L (lebesgue_on T) (\<lambda>x. if x \<in> S then f x else 0) = integral\<^sup>L (lebesgue_on S) f"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   455
  using integral_restrict_Int [of S T f] assms
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   456
  by (simp add: Int_absorb2)
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   457
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   458
lemma integral_restrict_UNIV:
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   459
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   460
  assumes "S \<in> sets lebesgue"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   461
  shows "integral\<^sup>L lebesgue (\<lambda>x. if x \<in> S then f x else 0) = integral\<^sup>L (lebesgue_on S) f"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   462
  using integral_restrict_Int [of S UNIV f] assms
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   463
  by (simp add: lebesgue_on_UNIV_eq)
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
   464
70688
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   465
lemma integrable_lebesgue_on_empty [iff]:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   466
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{second_countable_topology,banach}"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   467
  shows "integrable (lebesgue_on {}) f"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   468
  by (simp add: integrable_restrict_space)
70271
f7630118814c a few general lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   469
70688
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   470
lemma integral_lebesgue_on_empty [simp]:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   471
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{second_countable_topology,banach}"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   472
  shows "integral\<^sup>L (lebesgue_on {}) f = 0"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   473
  by (simp add: Bochner_Integration.integral_empty)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   474
lemma has_bochner_integral_restrict_space:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   475
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   476
  assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   477
  shows "has_bochner_integral (restrict_space M \<Omega>) f i
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   478
     \<longleftrightarrow> has_bochner_integral M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x) i"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   479
  by (simp add: integrable_restrict_space [OF assms] integral_restrict_space [OF assms] has_bochner_integral_iff)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   480
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   481
lemma integrable_restrict_UNIV:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   482
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   483
  assumes S: "S \<in> sets lebesgue"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   484
  shows  "integrable lebesgue (\<lambda>x. if x \<in> S then f x else 0) \<longleftrightarrow> integrable (lebesgue_on S) f"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   485
  using has_bochner_integral_restrict_space [of S lebesgue f] assms
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   486
  by (simp add: integrable.simps indicator_scaleR_eq_if)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   487
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   488
lemma integral_mono_lebesgue_on_AE:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   489
  fixes f::"_ \<Rightarrow> real"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   490
  assumes f: "integrable (lebesgue_on T) f"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   491
    and gf: "AE x in (lebesgue_on S). g x \<le> f x"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   492
    and f0: "AE x in (lebesgue_on T). 0 \<le> f x"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   493
    and "S \<subseteq> T" and S: "S \<in> sets lebesgue" and T: "T \<in> sets lebesgue"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   494
  shows "(\<integral>x. g x \<partial>(lebesgue_on S)) \<le> (\<integral>x. f x \<partial>(lebesgue_on T))"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   495
proof -
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   496
  have "(\<integral>x. g x \<partial>(lebesgue_on S)) = (\<integral>x. (if x \<in> S then g x else 0) \<partial>lebesgue)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   497
    by (simp add: Lebesgue_Measure.integral_restrict_UNIV S)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   498
  also have "\<dots> \<le> (\<integral>x. (if x \<in> T then f x else 0) \<partial>lebesgue)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   499
  proof (rule Bochner_Integration.integral_mono_AE')
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   500
    show "integrable lebesgue (\<lambda>x. if x \<in> T then f x else 0)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   501
      by (simp add: integrable_restrict_UNIV T f)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   502
    show "AE x in lebesgue. (if x \<in> S then g x else 0) \<le> (if x \<in> T then f x else 0)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   503
      using assms by (auto simp: AE_restrict_space_iff)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   504
    show "AE x in lebesgue. 0 \<le> (if x \<in> T then f x else 0)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   505
      using f0 by (simp add: AE_restrict_space_iff T)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   506
  qed
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   507
  also have "\<dots> = (\<integral>x. f x \<partial>(lebesgue_on T))"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   508
    using Lebesgue_Measure.integral_restrict_UNIV T by blast
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   509
  finally show ?thesis .
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   510
qed
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   511
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   512
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   513
subsection \<open>Borel measurability\<close>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   514
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   515
lemma borel_measurable_if_I:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   516
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   517
  assumes f: "f \<in> borel_measurable (lebesgue_on S)" and S: "S \<in> sets lebesgue"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   518
  shows "(\<lambda>x. if x \<in> S then f x else 0) \<in> borel_measurable lebesgue"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   519
proof -
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   520
  have eq: "{x. x \<notin> S} \<union> {x. f x \<in> Y} = {x. x \<notin> S} \<union> {x. f x \<in> Y} \<inter> S" for Y
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   521
    by blast
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   522
  show ?thesis
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   523
  using f S
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   524
  apply (simp add: vimage_def in_borel_measurable_borel Ball_def)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   525
  apply (elim all_forward imp_forward asm_rl)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   526
  apply (simp only: Collect_conj_eq Collect_disj_eq imp_conv_disj eq)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   527
  apply (auto simp: Compl_eq [symmetric] Compl_in_sets_lebesgue sets_restrict_space_iff)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   528
  done
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   529
qed
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   530
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   531
lemma borel_measurable_if_D:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   532
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   533
  assumes "(\<lambda>x. if x \<in> S then f x else 0) \<in> borel_measurable lebesgue"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   534
  shows "f \<in> borel_measurable (lebesgue_on S)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   535
  using assms
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   536
  apply (simp add: in_borel_measurable_borel Ball_def)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   537
  apply (elim all_forward imp_forward asm_rl)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   538
  apply (force simp: space_restrict_space sets_restrict_space image_iff intro: rev_bexI)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   539
  done
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   540
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   541
lemma borel_measurable_if:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   542
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   543
  assumes "S \<in> sets lebesgue"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   544
  shows "(\<lambda>x. if x \<in> S then f x else 0) \<in> borel_measurable lebesgue \<longleftrightarrow> f \<in> borel_measurable (lebesgue_on S)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   545
  using assms borel_measurable_if_D borel_measurable_if_I by blast
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   546
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   547
lemma borel_measurable_vimage_halfspace_component_lt:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   548
     "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   549
      (\<forall>a i. i \<in> Basis \<longrightarrow> {x \<in> S. f x \<bullet> i < a} \<in> sets (lebesgue_on S))"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   550
  apply (rule trans [OF borel_measurable_iff_halfspace_less])
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   551
  apply (fastforce simp add: space_restrict_space)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   552
  done
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   553
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   554
lemma borel_measurable_vimage_halfspace_component_ge:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   555
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   556
  shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   557
         (\<forall>a i. i \<in> Basis \<longrightarrow> {x \<in> S. f x \<bullet> i \<ge> a} \<in> sets (lebesgue_on S))"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   558
  apply (rule trans [OF borel_measurable_iff_halfspace_ge])
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   559
  apply (fastforce simp add: space_restrict_space)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   560
  done
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   561
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   562
lemma borel_measurable_vimage_halfspace_component_gt:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   563
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   564
  shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   565
         (\<forall>a i. i \<in> Basis \<longrightarrow> {x \<in> S. f x \<bullet> i > a} \<in> sets (lebesgue_on S))"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   566
  apply (rule trans [OF borel_measurable_iff_halfspace_greater])
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   567
  apply (fastforce simp add: space_restrict_space)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   568
  done
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   569
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   570
lemma borel_measurable_vimage_halfspace_component_le:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   571
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   572
  shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   573
         (\<forall>a i. i \<in> Basis \<longrightarrow> {x \<in> S. f x \<bullet> i \<le> a} \<in> sets (lebesgue_on S))"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   574
  apply (rule trans [OF borel_measurable_iff_halfspace_le])
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   575
  apply (fastforce simp add: space_restrict_space)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   576
  done
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   577
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   578
lemma
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   579
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   580
  shows borel_measurable_vimage_open_interval:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   581
         "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   582
         (\<forall>a b. {x \<in> S. f x \<in> box a b} \<in> sets (lebesgue_on S))" (is ?thesis1)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   583
   and borel_measurable_vimage_open:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   584
         "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   585
         (\<forall>T. open T \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets (lebesgue_on S))" (is ?thesis2)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   586
proof -
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   587
  have "{x \<in> S. f x \<in> box a b} \<in> sets (lebesgue_on S)" if "f \<in> borel_measurable (lebesgue_on S)" for a b
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   588
  proof -
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   589
    have "S = S \<inter> space lebesgue"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   590
      by simp
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   591
    then have "S \<inter> (f -` box a b) \<in> sets (lebesgue_on S)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   592
      by (metis (no_types) box_borel in_borel_measurable_borel inf_sup_aci(1) space_restrict_space that)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   593
    then show ?thesis
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   594
      by (simp add: Collect_conj_eq vimage_def)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   595
  qed
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   596
  moreover
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   597
  have "{x \<in> S. f x \<in> T} \<in> sets (lebesgue_on S)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   598
       if T: "\<And>a b. {x \<in> S. f x \<in> box a b} \<in> sets (lebesgue_on S)" "open T" for T
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   599
  proof -
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   600
    obtain \<D> where "countable \<D>" and \<D>: "\<And>X. X \<in> \<D> \<Longrightarrow> \<exists>a b. X = box a b" "\<Union>\<D> = T"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   601
      using open_countable_Union_open_box that \<open>open T\<close> by metis
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   602
    then have eq: "{x \<in> S. f x \<in> T} = (\<Union>U \<in> \<D>. {x \<in> S. f x \<in> U})"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   603
      by blast
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   604
    have "{x \<in> S. f x \<in> U} \<in> sets (lebesgue_on S)" if "U \<in> \<D>" for U
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   605
      using that T \<D> by blast
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   606
    then show ?thesis
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   607
      by (auto simp: eq intro: Sigma_Algebra.sets.countable_UN' [OF \<open>countable \<D>\<close>])
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   608
  qed
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   609
  moreover
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   610
  have eq: "{x \<in> S. f x \<bullet> i < a} = {x \<in> S. f x \<in> {y. y \<bullet> i < a}}" for i a
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   611
    by auto
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   612
  have "f \<in> borel_measurable (lebesgue_on S)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   613
    if "\<And>T. open T \<Longrightarrow> {x \<in> S. f x \<in> T} \<in> sets (lebesgue_on S)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   614
    by (metis (no_types) eq borel_measurable_vimage_halfspace_component_lt open_halfspace_component_lt that)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   615
  ultimately show "?thesis1" "?thesis2"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   616
    by blast+
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   617
qed
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   618
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   619
lemma borel_measurable_vimage_closed:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   620
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   621
  shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   622
         (\<forall>T. closed T \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets (lebesgue_on S))"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   623
        (is "?lhs = ?rhs")
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   624
proof -
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   625
  have eq: "{x \<in> S. f x \<in> T} = S - {x \<in> S. f x \<in> (- T)}" for T
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   626
    by auto
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   627
  show ?thesis
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   628
    apply (simp add: borel_measurable_vimage_open, safe)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   629
     apply (simp_all (no_asm) add: eq)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   630
     apply (intro sets.Diff sets_lebesgue_on_refl, force simp: closed_open)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   631
    apply (intro sets.Diff sets_lebesgue_on_refl, force simp: open_closed)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   632
    done
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   633
qed
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   634
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   635
lemma borel_measurable_vimage_closed_interval:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   636
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   637
  shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   638
         (\<forall>a b. {x \<in> S. f x \<in> cbox a b} \<in> sets (lebesgue_on S))"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   639
        (is "?lhs = ?rhs")
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   640
proof
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   641
  assume ?lhs then show ?rhs
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   642
    using borel_measurable_vimage_closed by blast
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   643
next
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   644
  assume RHS: ?rhs
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   645
  have "{x \<in> S. f x \<in> T} \<in> sets (lebesgue_on S)" if "open T" for T
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   646
  proof -
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   647
    obtain \<D> where "countable \<D>" and \<D>: "\<D> \<subseteq> Pow T" "\<And>X. X \<in> \<D> \<Longrightarrow> \<exists>a b. X = cbox a b" "\<Union>\<D> = T"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   648
      using open_countable_Union_open_cbox that \<open>open T\<close> by metis
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   649
    then have eq: "{x \<in> S. f x \<in> T} = (\<Union>U \<in> \<D>. {x \<in> S. f x \<in> U})"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   650
      by blast
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   651
    have "{x \<in> S. f x \<in> U} \<in> sets (lebesgue_on S)" if "U \<in> \<D>" for U
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   652
      using that \<D> by (metis RHS)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   653
    then show ?thesis
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   654
      by (auto simp: eq intro: Sigma_Algebra.sets.countable_UN' [OF \<open>countable \<D>\<close>])
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   655
  qed
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   656
  then show ?lhs
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   657
    by (simp add: borel_measurable_vimage_open)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   658
qed
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   659
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   660
lemma borel_measurable_UNIV_eq: "borel_measurable (lebesgue_on UNIV) = borel_measurable lebesgue"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   661
  by auto
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   662
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   663
lemma borel_measurable_vimage_borel:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   664
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   665
  shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   666
         (\<forall>T. T \<in> sets borel \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets (lebesgue_on S))"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   667
        (is "?lhs = ?rhs")
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   668
proof
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   669
  assume f: ?lhs
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   670
  then show ?rhs
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   671
    using measurable_sets [OF f]
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   672
      by (simp add: Collect_conj_eq inf_sup_aci(1) space_restrict_space vimage_def)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   673
qed (simp add: borel_measurable_vimage_open_interval)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   674
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   675
lemma lebesgue_measurable_vimage_borel:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   676
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   677
  assumes "f \<in> borel_measurable lebesgue" "T \<in> sets borel"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   678
  shows "{x. f x \<in> T} \<in> sets lebesgue"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   679
  using assms borel_measurable_vimage_borel [of f UNIV] by auto
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   680
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   681
lemma borel_measurable_lebesgue_preimage_borel:
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   682
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   683
  shows "f \<in> borel_measurable lebesgue \<longleftrightarrow>
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   684
         (\<forall>T. T \<in> sets borel \<longrightarrow> {x. f x \<in> T} \<in> sets lebesgue)"
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   685
  apply (intro iffI allI impI lebesgue_measurable_vimage_borel)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   686
    apply (auto simp: in_borel_measurable_borel vimage_def)
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   687
  done
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   688
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   689
3d894e1cfc75 new material on Analysis, plus some rearrangements
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
   690
subsection \<^marker>\<open>tag unimportant\<close> \<open>Measurability of continuous functions\<close>
69447
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
   691
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   692
lemma continuous_imp_measurable_on_sets_lebesgue:
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   693
  assumes f: "continuous_on S f" and S: "S \<in> sets lebesgue"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   694
  shows "f \<in> borel_measurable (lebesgue_on S)"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   695
proof -
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   696
  have "sets (restrict_space borel S) \<subseteq> sets (lebesgue_on S)"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   697
    by (simp add: mono_restrict_space subsetI)
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   698
  then show ?thesis
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   699
    by (simp add: borel_measurable_continuous_on_restrict [OF f] borel_measurable_subalgebra 
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   700
                  space_restrict_space)
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   701
qed
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   702
70380
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   703
lemma id_borel_measurable_lebesgue [iff]: "id \<in> borel_measurable lebesgue"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   704
  by (simp add: measurable_completion)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   705
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   706
lemma id_borel_measurable_lebesgue_on [iff]: "id \<in> borel_measurable (lebesgue_on S)"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   707
  by (simp add: measurable_completion measurable_restrict_space1)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   708
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   709
context
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   710
begin
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   711
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 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
   713
  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
   714
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
   715
  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
   716
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   717
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
   718
  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
   719
  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
   720
  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
   721
     (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
   722
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   723
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
   724
  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
   725
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   726
lemma nn_integral_lborel_prod:
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   727
  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
   728
  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
   729
  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))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   730
  by (simp add: lborel_def nn_integral_distr product_nn_integral_prod
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   731
                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
   732
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   733
lemma emeasure_lborel_Icc[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
   734
  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
   735
  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
   736
  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
   737
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
   738
  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
   739
    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
   740
  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
   741
    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
   742
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   743
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   744
lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ennreal (if l \<le> u then u - l else 0)"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   745
  by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   746
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   747
lemma\<^marker>\<open>tag important\<close> emeasure_lborel_cbox[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
   748
  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
   749
  shows "emeasure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   750
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   751
  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b) :: ennreal) = indicator (cbox l u)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   752
    by (auto simp: fun_eq_iff cbox_def split: split_indicator)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   753
  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
   754
    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
   755
  also have "\<dots> = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   756
    by (subst nn_integral_lborel_prod) (simp_all add: prod_ennreal inner_diff_left)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   757
  finally show ?thesis .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   758
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   759
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   760
lemma AE_lborel_singleton: "AE x in lborel::'a::euclidean_space measure. x \<noteq> c"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   761
  using SOME_Basis AE_discrete_difference [of "{c}" lborel] emeasure_lborel_cbox [of c c]
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   762
  by (auto simp add: power_0_left)
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   763
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   764
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
   765
  assumes [simp]: "l \<le> u"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   766
  shows "emeasure lborel {l <..< u} = ennreal (u - l)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   767
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
   768
  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
   769
    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
   770
  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
   771
    by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   772
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   773
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   774
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
   775
  assumes [simp]: "l \<le> u"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   776
  shows "emeasure lborel {l <.. u} = ennreal (u - l)"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   777
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
   778
  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
   779
    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
   780
  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
   781
    by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   782
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   783
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   784
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
   785
  assumes [simp]: "l \<le> u"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   786
  shows "emeasure lborel {l ..< u} = ennreal (u - l)"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   787
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   788
  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
   789
    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
   790
  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
   791
    by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   792
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   793
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   794
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
   795
  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
   796
  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
   797
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   798
  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b <..< u\<bullet>b} (x \<bullet> b) :: ennreal) = indicator (box l u)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   799
    by (auto simp: fun_eq_iff box_def split: split_indicator)
57447
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 "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
   801
    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
   802
  also have "\<dots> = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   803
    by (subst nn_integral_lborel_prod) (simp_all add: prod_ennreal inner_diff_left)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   804
  finally show ?thesis .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   805
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   806
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   807
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
   808
  "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
   809
  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
   810
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   811
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
   812
  "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
   813
  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
   814
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   815
lemma emeasure_lborel_singleton[simp]: "emeasure lborel {x} = 0"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   816
  using emeasure_lborel_cbox[of x x] nonempty_Basis
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   817
  by (auto simp del: emeasure_lborel_cbox nonempty_Basis simp add: prod_constant)
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   818
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   819
lemma fmeasurable_cbox [iff]: "cbox a b \<in> fmeasurable lborel"
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   820
  and fmeasurable_box [iff]: "box a b \<in> fmeasurable lborel"
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 65680
diff changeset
   821
  by (auto simp: fmeasurable_def emeasure_lborel_box_eq emeasure_lborel_cbox_eq)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   822
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   823
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
   824
  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
   825
  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
   826
  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
   827
    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
   828
    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
   829
    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
   830
  by (simp_all add: measure_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   831
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   832
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
   833
  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
   834
  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
   835
    and measure_lborel_cbox[simp]: "measure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   836
  by (simp_all add: measure_def inner_diff_left prod_nonneg)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   837
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   838
lemma measure_lborel_cbox_eq:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   839
  "measure 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)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   840
  using box_eq_empty(2)[THEN iffD2, of u l] by (auto simp: not_le)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   841
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   842
lemma measure_lborel_box_eq:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   843
  "measure 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)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   844
  using box_eq_empty(1)[THEN iffD2, of u l] by (auto simp: not_le dest!: less_imp_le) force
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   845
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   846
lemma measure_lborel_singleton[simp]: "measure lborel {x} = 0"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   847
  by (simp add: measure_def)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   848
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   849
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
   850
proof
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   851
  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
   852
    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
   853
       (auto simp: emeasure_lborel_cbox_eq UN_box_eq_UNIV)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   854
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   855
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   856
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
   857
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   858
lemma emeasure_lborel_UNIV [simp]: "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
   859
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
   860
  { 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
   861
    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
   862
    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
   863
      by (simp add: mult_le_cancel_right1)
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   864
    also
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
   865
    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
   866
      apply (rule mult_left_mono)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
   867
      apply (metis DIM_positive One_nat_def less_eq_Suc_le less_imp_le of_nat_le_iff of_nat_power self_le_power zero_less_Suc)
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
   868
      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
   869
      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
   870
    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
   871
  } 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
   872
  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
   873
    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
   874
    apply (subst SUP_emeasure_incseq[symmetric])
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   875
    apply (auto simp: incseq_def subset_box inner_add_left prod_constant
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   876
      simp del: Sup_eq_top_iff SUP_eq_top_iff
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   877
      intro!: ennreal_SUP_eq_top)
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   878
    done
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
   879
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   880
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   881
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
   882
  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
   883
  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
   884
  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
   885
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   886
  have "A \<subseteq> (\<Union>i. {from_nat_into A i})" using from_nat_into_surj assms by force
63262
e497387de7af remove smt call in Lebesge_Measure
hoelzl
parents: 63040
diff changeset
   887
  then have "emeasure lborel A \<le> emeasure lborel (\<Union>i. {from_nat_into A i})"
e497387de7af remove smt call in Lebesge_Measure
hoelzl
parents: 63040
diff changeset
   888
    by (intro emeasure_mono) auto
e497387de7af remove smt call in Lebesge_Measure
hoelzl
parents: 63040
diff changeset
   889
  also have "emeasure lborel (\<Union>i. {from_nat_into A i}) = 0"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   890
    by (rule emeasure_UN_eq_0) auto
63262
e497387de7af remove smt call in Lebesge_Measure
hoelzl
parents: 63040
diff changeset
   891
  finally show ?thesis
e497387de7af remove smt call in Lebesge_Measure
hoelzl
parents: 63040
diff changeset
   892
    by (auto simp add: )
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   893
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   894
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   895
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
   896
  by (simp add: null_sets_def emeasure_lborel_countable sets.countable)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   897
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   898
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
   899
  by (intro countable_imp_null_set_lborel countable_finite)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   900
70380
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   901
lemma insert_null_sets_iff [simp]: "insert a N \<in> null_sets lebesgue \<longleftrightarrow> N \<in> null_sets lebesgue"     
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   902
    (is "?lhs = ?rhs")
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   903
proof
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   904
  assume ?lhs then show ?rhs
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   905
    by (meson completion.complete2 subset_insertI)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   906
next
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   907
  assume ?rhs then show ?lhs
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   908
  by (simp add: null_sets.insert_in_sets null_setsI)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   909
qed
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   910
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   911
lemma insert_null_sets_lebesgue_on_iff [simp]:
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   912
  assumes "a \<in> S" "S \<in> sets lebesgue"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   913
  shows "insert a N \<in> null_sets (lebesgue_on S) \<longleftrightarrow> N \<in> null_sets (lebesgue_on S)"     
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   914
  by (simp add: assms null_sets_restrict_space)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
   915
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   916
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
   917
proof
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   918
  assume asm: "lborel = count_space A"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   919
  have "space lborel = UNIV" by simp
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   920
  hence [simp]: "A = UNIV" by (subst (asm) asm) (simp only: space_count_space)
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
   921
  have "emeasure lborel {undefined::'a} = 1"
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   922
      by (subst asm, subst emeasure_count_space_finite) auto
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   923
  moreover have "emeasure lborel {undefined} \<noteq> 1" by simp
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   924
  ultimately show False by contradiction
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   925
qed
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
   926
65204
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   927
lemma mem_closed_if_AE_lebesgue_open:
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   928
  assumes "open S" "closed C"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   929
  assumes "AE x \<in> S in lebesgue. x \<in> C"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   930
  assumes "x \<in> S"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   931
  shows "x \<in> C"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   932
proof (rule ccontr)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   933
  assume xC: "x \<notin> C"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   934
  with openE[of "S - C"] assms
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   935
  obtain e where e: "0 < e" "ball x e \<subseteq> S - C"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   936
    by blast
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   937
  then obtain a b where box: "x \<in> box a b" "box a b \<subseteq> S - C"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   938
    by (metis rational_boxes order_trans)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   939
  then have "0 < emeasure lebesgue (box a b)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   940
    by (auto simp: emeasure_lborel_box_eq mem_box algebra_simps intro!: prod_pos)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   941
  also have "\<dots> \<le> emeasure lebesgue (S - C)"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   942
    using assms box
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   943
    by (auto intro!: emeasure_mono)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   944
  also have "\<dots> = 0"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   945
    using assms
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   946
    by (auto simp: eventually_ae_filter completion.complete2 set_diff_eq null_setsD1)
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   947
  finally show False by simp
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   948
qed
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   949
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   950
lemma mem_closed_if_AE_lebesgue: "closed C \<Longrightarrow> (AE x in lebesgue. x \<in> C) \<Longrightarrow> x \<in> C"
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   951
  using mem_closed_if_AE_lebesgue_open[OF open_UNIV] by simp
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   952
d23eded35a33 modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents: 64272
diff changeset
   953
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   954
subsection \<open>Affine transformation on the Lebesgue-Borel\<close>
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   955
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   956
lemma\<^marker>\<open>tag important\<close> 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
   957
  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
   958
  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
   959
  assumes sets_eq: "sets M = sets borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   960
  shows "lborel = M"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   961
proof (rule measure_eqI_generator_eq)
57447
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 ?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
   963
  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
   964
    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
   965
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   966
  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
   967
    by (simp_all add: borel_eq_box sets_eq)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   968
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   969
  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
   970
  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
   971
    unfolding UN_box_eq_UNIV by auto
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   972
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   973
  { fix i show "emeasure lborel (?A i) \<noteq> \<infinity>" by auto }
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   974
  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   975
      apply (auto simp: emeasure_eq emeasure_lborel_box_eq)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   976
      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
   977
      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
   978
      done }
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   979
qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   980
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   981
lemma\<^marker>\<open>tag important\<close> lborel_affine_euclidean:
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   982
  fixes c :: "'a::euclidean_space \<Rightarrow> real" and t
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   983
  defines "T x \<equiv> t + (\<Sum>j\<in>Basis. (c j * (x \<bullet> j)) *\<^sub>R j)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   984
  assumes c: "\<And>j. j \<in> Basis \<Longrightarrow> c j \<noteq> 0"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   985
  shows "lborel = density (distr lborel borel T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))" (is "_ = ?D")
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   986
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
   987
  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
   988
  fix l u assume le: "\<And>b. b \<in> ?B \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   989
  have [measurable]: "T \<in> borel \<rightarrow>\<^sub>M borel"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   990
    by (simp add: T_def[abs_def])
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   991
  have eq: "T -` box l u = box
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   992
    (\<Sum>j\<in>Basis. (((if 0 < c j then l - t else u - t) \<bullet> j) / c j) *\<^sub>R j)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   993
    (\<Sum>j\<in>Basis. (((if 0 < c j then u - t else l - t) \<bullet> j) / c j) *\<^sub>R j)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   994
    using c by (auto simp: box_def T_def field_simps inner_simps divide_less_eq)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   995
  with le c show "emeasure ?D (box l u) = (\<Prod>b\<in>?B. (u - l) \<bullet> b)"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   996
    by (auto simp: emeasure_density emeasure_distr nn_integral_multc emeasure_lborel_box_eq inner_simps
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   997
                   field_simps divide_simps ennreal_mult'[symmetric] prod_nonneg prod.distrib[symmetric]
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   998
             intro!: prod.cong)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   999
qed simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1000
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
  1001
lemma lborel_affine:
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
  1002
  fixes t :: "'a::euclidean_space"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
  1003
  shows "c \<noteq> 0 \<Longrightarrow> lborel = density (distr lborel borel (\<lambda>x. t + c *\<^sub>R x)) (\<lambda>_. \<bar>c\<bar>^DIM('a))"
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
  1004
  using lborel_affine_euclidean[where c="\<lambda>_::'a. c" and t=t]
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1005
  unfolding scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] euclidean_representation prod_constant by simp
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
  1006
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1007
lemma lborel_real_affine:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1008
  "c \<noteq> 0 \<Longrightarrow> lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. ennreal (abs 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
  1009
  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
  1010
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
  1011
lemma AE_borel_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
  1012
  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
  1013
  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
  1014
  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
  1015
     (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
  1016
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1017
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
  1018
  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
  1019
  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
  1020
  by (subst lborel_real_affine[OF c, of t])
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1021
     (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
  1022
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1023
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
  1024
  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
  1025
  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
  1026
  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
  1027
  using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1028
  by (subst (asm) nn_integral_real_affine[where c=c and t=t]) (auto simp: ennreal_mult_less_top)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1029
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1030
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
  1031
  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
  1032
  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
  1033
  using
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1034
    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
  1035
    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
  1036
  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
  1037
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1038
lemma\<^marker>\<open>tag important\<close> lborel_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
  1039
  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
  1040
  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)"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1041
proof cases
57166
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57138
diff changeset
  1042
  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
  1043
    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
  1044
    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
  1045
       (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
  1046
next
5cfcc616d485 use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents: 57138
diff changeset
  1047
  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
  1048
    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
  1049
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
  1050
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1051
lemma
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1052
  fixes c :: "'a::euclidean_space \<Rightarrow> real" and t
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1053
  assumes c: "\<And>j. j \<in> Basis \<Longrightarrow> c j \<noteq> 0"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1054
  defines "T == (\<lambda>x. t + (\<Sum>j\<in>Basis. (c j * (x \<bullet> j)) *\<^sub>R j))"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1055
  shows lebesgue_affine_euclidean: "lebesgue = density (distr lebesgue lebesgue T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))" (is "_ = ?D")
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1056
    and lebesgue_affine_measurable: "T \<in> lebesgue \<rightarrow>\<^sub>M lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1057
proof -
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1058
  have T_borel[measurable]: "T \<in> borel \<rightarrow>\<^sub>M borel"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1059
    by (auto simp: T_def[abs_def])
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1060
  { fix A :: "'a set" assume A: "A \<in> sets borel"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1061
    then have "emeasure lborel A = 0 \<longleftrightarrow> emeasure (density (distr lborel borel T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))) A = 0"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1062
      unfolding T_def using c by (subst lborel_affine_euclidean[symmetric]) auto
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1063
    also have "\<dots> \<longleftrightarrow> emeasure (distr lebesgue lborel T) A = 0"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1064
      using A c by (simp add: distr_completion emeasure_density nn_integral_cmult prod_nonneg cong: distr_cong)
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1065
    finally have "emeasure lborel A = 0 \<longleftrightarrow> emeasure (distr lebesgue lborel T) A = 0" . }
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1066
  then have eq: "null_sets lborel = null_sets (distr lebesgue lborel T)"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1067
    by (auto simp: null_sets_def)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1068
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1069
  show "T \<in> lebesgue \<rightarrow>\<^sub>M lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1070
    by (rule completion.measurable_completion2) (auto simp: eq measurable_completion)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1071
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1072
  have "lebesgue = completion (density (distr lborel borel T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>)))"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1073
    using c by (subst lborel_affine_euclidean[of c t]) (simp_all add: T_def[abs_def])
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1074
  also have "\<dots> = density (completion (distr lebesgue lborel T)) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1075
    using c by (auto intro!: always_eventually prod_pos completion_density_eq simp: distr_completion cong: distr_cong)
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1076
  also have "\<dots> = density (distr lebesgue lebesgue T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1077
    by (subst completion.completion_distr_eq) (auto simp: eq measurable_completion)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1078
  finally show "lebesgue = density (distr lebesgue lebesgue T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))" .
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1079
qed
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1080
70547
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1081
corollary lebesgue_real_affine:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1082
  "c \<noteq> 0 \<Longrightarrow> lebesgue = density (distr lebesgue lebesgue (\<lambda>x. t + c * x)) (\<lambda>_. ennreal (abs c))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1083
    using lebesgue_affine_euclidean [where c= "\<lambda>x::real. c"] by simp
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1084
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1085
lemma nn_integral_real_affine_lebesgue:
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1086
  fixes c :: real assumes f[measurable]: "f \<in> borel_measurable lebesgue" and c: "c \<noteq> 0"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1087
  shows "(\<integral>\<^sup>+x. f x \<partial>lebesgue) = ennreal\<bar>c\<bar> * (\<integral>\<^sup>+x. f(t + c * x) \<partial>lebesgue)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1088
proof -
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1089
  have "(\<integral>\<^sup>+x. f x \<partial>lebesgue) = (\<integral>\<^sup>+x. f x \<partial>density (distr lebesgue lebesgue (\<lambda>x. t + c * x)) (\<lambda>x. ennreal \<bar>c\<bar>))"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1090
    using lebesgue_real_affine c by auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1091
  also have "\<dots> = \<integral>\<^sup>+ x. ennreal \<bar>c\<bar> * f x \<partial>distr lebesgue lebesgue (\<lambda>x. t + c * x)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1092
    by (subst nn_integral_density) auto
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1093
  also have "\<dots> = ennreal \<bar>c\<bar> * integral\<^sup>N (distr lebesgue lebesgue (\<lambda>x. t + c * x)) f"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1094
    using f measurable_distr_eq1 nn_integral_cmult by blast
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1095
  also have "\<dots> = \<bar>c\<bar> * (\<integral>\<^sup>+x. f(t + c * x) \<partial>lebesgue)"
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1096
    using lebesgue_affine_measurable[where c= "\<lambda>x::real. c"]
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1097
    by (subst nn_integral_distr) (force+)
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1098
  finally show ?thesis .
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1099
qed
7ce95a5c4aa8 new material on eqiintegrable functions, etc.
paulson <lp15@cam.ac.uk>
parents: 70532
diff changeset
  1100
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1101
lemma lebesgue_measurable_scaling[measurable]: "(*\<^sub>R) x \<in> lebesgue \<rightarrow>\<^sub>M lebesgue"
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1102
proof cases
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1103
  assume "x = 0"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1104
  then have "(*\<^sub>R) x = (\<lambda>x. 0::'a)"
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1105
    by (auto simp: fun_eq_iff)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1106
  then show ?thesis by auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1107
next
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1108
  assume "x \<noteq> 0" then show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1109
    using lebesgue_affine_measurable[of "\<lambda>_. x" 0]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1110
    unfolding scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] euclidean_representation
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1111
    by (auto simp add: ac_simps)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1112
qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1113
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1114
lemma
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1115
  fixes m :: real and \<delta> :: "'a::euclidean_space"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1116
  defines "T r d x \<equiv> r *\<^sub>R x + d"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1117
  shows emeasure_lebesgue_affine: "emeasure lebesgue (T m \<delta> ` S) = \<bar>m\<bar> ^ DIM('a) * emeasure lebesgue S" (is ?e)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1118
    and measure_lebesgue_affine: "measure lebesgue (T m \<delta> ` S) = \<bar>m\<bar> ^ DIM('a) * measure lebesgue S" (is ?m)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1119
proof -
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1120
  show ?e
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1121
  proof cases
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1122
    assume "m = 0" then show ?thesis
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1123
      by (simp add: image_constant_conv T_def[abs_def])
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1124
  next
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1125
    let ?T = "T m \<delta>" and ?T' = "T (1 / m) (- ((1/m) *\<^sub>R \<delta>))"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1126
    assume "m \<noteq> 0"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1127
    then have s_comp_s: "?T' \<circ> ?T = id" "?T \<circ> ?T' = id"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1128
      by (auto simp: T_def[abs_def] fun_eq_iff scaleR_add_right scaleR_diff_right)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1129
    then have "inv ?T' = ?T" "bij ?T'"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1130
      by (auto intro: inv_unique_comp o_bij)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1131
    then have eq: "T m \<delta> ` S = T (1 / m) ((-1/m) *\<^sub>R \<delta>) -` S \<inter> space lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1132
      using bij_vimage_eq_inv_image[OF \<open>bij ?T'\<close>, of S] by auto
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1133
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1134
    have trans_eq_T: "(\<lambda>x. \<delta> + (\<Sum>j\<in>Basis. (m * (x \<bullet> j)) *\<^sub>R j)) = T m \<delta>" for m \<delta>
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1135
      unfolding T_def[abs_def] scaleR_scaleR[symmetric] scaleR_sum_right[symmetric]
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1136
      by (auto simp add: euclidean_representation ac_simps)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1137
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1138
    have T[measurable]: "T r d \<in> lebesgue \<rightarrow>\<^sub>M lebesgue" for r d
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1139
      using lebesgue_affine_measurable[of "\<lambda>_. r" d]
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1140
      by (cases "r = 0") (auto simp: trans_eq_T T_def[abs_def])
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1141
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1142
    show ?thesis
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1143
    proof cases
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1144
      assume "S \<in> sets lebesgue" with \<open>m \<noteq> 0\<close> show ?thesis
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1145
        unfolding eq
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1146
        apply (subst lebesgue_affine_euclidean[of "\<lambda>_. m" \<delta>])
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1147
        apply (simp_all add: emeasure_density trans_eq_T nn_integral_cmult emeasure_distr
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1148
                        del: space_completion emeasure_completion)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1149
        apply (simp add: vimage_comp s_comp_s prod_constant)
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1150
        done
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1151
    next
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1152
      assume "S \<notin> sets lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1153
      moreover have "?T ` S \<notin> sets lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1154
      proof
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1155
        assume "?T ` S \<in> sets lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1156
        then have "?T -` (?T ` S) \<inter> space lebesgue \<in> sets lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1157
          by (rule measurable_sets[OF T])
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1158
        also have "?T -` (?T ` S) \<inter> space lebesgue = S"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1159
          by (simp add: vimage_comp s_comp_s eq)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1160
        finally show False using \<open>S \<notin> sets lebesgue\<close> by auto
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1161
      qed
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1162
      ultimately show ?thesis
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1163
        by (simp add: emeasure_notin_sets)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1164
    qed
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1165
  qed
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1166
  show ?m
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1167
    unfolding measure_def \<open>?e\<close> by (simp add: enn2real_mult prod_nonneg)
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1168
qed
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1169
67135
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66164
diff changeset
  1170
lemma lebesgue_real_scale:
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66164
diff changeset
  1171
  assumes "c \<noteq> 0"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66164
diff changeset
  1172
  shows   "lebesgue = density (distr lebesgue lebesgue (\<lambda>x. c * x)) (\<lambda>x. ennreal \<bar>c\<bar>)"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66164
diff changeset
  1173
  using assms by (subst lebesgue_affine_euclidean[of "\<lambda>_. c" 0]) simp_all
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66164
diff changeset
  1174
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
  1175
lemma divideR_right:
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
  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
  1177
  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
  1178
  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
  1179
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1180
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
  1181
  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
  1182
  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
  1183
    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
  1184
    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
  1185
  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
  1186
  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
  1187
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1188
lemma lborel_distr_uminus: "distr lborel borel uminus = (lborel :: real measure)"
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
  1189
  by (subst lborel_real_affine[of "-1" 0])
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1190
     (auto simp: density_1 one_ennreal_def[symmetric])
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1191
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
  1192
lemma lborel_distr_mult:
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1193
  assumes "(c::real) \<noteq> 0"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1194
  shows "distr lborel borel ((*) c) = density lborel (\<lambda>_. inverse \<bar>c\<bar>)"
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1195
proof-
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1196
  have "distr lborel borel ((*) c) = distr lborel lborel ((*) c)" by (simp cong: distr_cong)
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1197
  also from assms have "... = density lborel (\<lambda>_. inverse \<bar>c\<bar>)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1198
    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
  1199
  finally show ?thesis .
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1200
qed
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1201
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59741
diff changeset
  1202
lemma lborel_distr_mult':
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1203
  assumes "(c::real) \<noteq> 0"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1204
  shows "lborel = density (distr lborel borel ((*) c)) (\<lambda>_. \<bar>c\<bar>)"
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1205
proof-
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1206
  have "lborel = density lborel (\<lambda>_. 1)" by (rule density_1[symmetric])
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1207
  also from assms have "(\<lambda>_. 1 :: ennreal) = (\<lambda>_. inverse \<bar>c\<bar> * \<bar>c\<bar>)" by (intro ext) simp
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61808
diff changeset
  1208
  also have "density lborel ... = density (density lborel (\<lambda>_. inverse \<bar>c\<bar>)) (\<lambda>_. \<bar>c\<bar>)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1209
    by (subst density_density_eq) (auto simp: ennreal_mult)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1210
  also from assms have "density lborel (\<lambda>_. inverse \<bar>c\<bar>) = distr lborel borel ((*) c)"
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1211
    by (rule lborel_distr_mult[symmetric])
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1212
  finally show ?thesis .
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1213
qed
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1214
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67135
diff changeset
  1215
lemma lborel_distr_plus: "distr lborel borel ((+) c) = (lborel :: real measure)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1216
  by (subst lborel_real_affine[of 1 c]) (auto simp: density_1 one_ennreal_def[symmetric])
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1217
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61284
diff changeset
  1218
interpretation lborel: sigma_finite_measure lborel
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1219
  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
  1220
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1221
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
  1222
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1223
lemma\<^marker>\<open>tag important\<close> lborel_prod:
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1224
  "lborel \<Otimes>\<^sub>M lborel = (lborel :: ('a::euclidean_space \<times> 'b::euclidean_space) measure)"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1225
proof (rule lborel_eqI[symmetric], clarify)
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1226
  fix la ua :: 'a and lb ub :: 'b
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1227
  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
  1228
  have [simp]:
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1229
    "\<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
  1230
    "\<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
  1231
    "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
  1232
    "(\<lambda>u. (u, 0)) ` Basis \<inter> (\<lambda>u. (0, u)) ` Basis = {}"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1233
    "box (la, lb) (ua, ub) = box la ua \<times> box lb ub"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1234
    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
  1235
  show "emeasure (lborel \<Otimes>\<^sub>M lborel) (box (la, lb) (ua, ub)) =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67135
diff changeset
  1236
      ennreal (prod ((\<bullet>) ((ua, ub) - (la, lb))) Basis)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1237
    by (simp add: lborel.emeasure_pair_measure_Times Basis_prod_def prod.union_disjoint
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1238
                  prod.reindex ennreal_mult inner_diff_left prod_nonneg)
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1239
qed (simp add: borel_prod[symmetric])
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59357
diff changeset
  1240
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1241
(* FIXME: conversion in measurable prover *)
68120
2f161c6910f7 tidying more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1242
lemma lborelD_Collect[measurable (raw)]: "{x\<in>space borel. P x} \<in> sets borel \<Longrightarrow> {x\<in>space lborel. P x} \<in> sets lborel" 
2f161c6910f7 tidying more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1243
  by simp
2f161c6910f7 tidying more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1244
2f161c6910f7 tidying more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1245
lemma lborelD[measurable (raw)]: "A \<in> sets borel \<Longrightarrow> A \<in> sets lborel"
2f161c6910f7 tidying more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1246
  by 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
  1247
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1248
lemma emeasure_bounded_finite:
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1249
  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
  1250
proof -
68120
2f161c6910f7 tidying more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1251
  obtain a b where "A \<subseteq> cbox a b"
2f161c6910f7 tidying more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1252
    by (meson bounded_subset_cbox_symmetric \<open>bounded A\<close>)
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1253
  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
  1254
    by (intro emeasure_mono) auto
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1255
  then show ?thesis
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1256
    by (auto simp: emeasure_lborel_cbox_eq prod_nonneg less_top[symmetric] top_unique split: if_split_asm)
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1257
qed
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1258
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1259
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
  1260
  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
  1261
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1262
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
  1263
  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
  1264
  assumes "compact S" "continuous_on S f"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1265
  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
  1266
proof cases
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1267
  assume "S \<noteq> {}"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1268
  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
  1269
    using assms by (intro continuous_intros)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
  1270
  from continuous_attains_sup[OF \<open>compact S\<close> \<open>S \<noteq> {}\<close> this]
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1271
  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
  1272
    by auto
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1273
  show ?thesis
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1274
  proof (rule integrable_bound)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1275
    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
  1276
      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
  1277
    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
  1278
      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
  1279
    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
  1280
      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
  1281
  qed
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1282
qed simp
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1283
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1284
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
  1285
  fixes f :: "real \<Rightarrow> real"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1286
  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
  1287
  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
  1288
proof -
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1289
  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
  1290
  proof (rule borel_integrable_compact)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1291
    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
  1292
      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
  1293
  qed simp
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1294
  then show ?thesis
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57447
diff changeset
  1295
    by (auto simp: mult.commute)
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
  1296
qed
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1297
69447
b7b9cbe0bd43 Tagged some of HOL-Analysis
eberlm <eberlm@in.tum.de>
parents: 69260
diff changeset
  1298
subsection \<open>Lebesgue measurable sets\<close>
67984
adc1a992c470 a few more results
paulson <lp15@cam.ac.uk>
parents: 67982
diff changeset
  1299
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1300
abbreviation\<^marker>\<open>tag important\<close> lmeasurable :: "'a::euclidean_space set set"
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1301
where
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1302
  "lmeasurable \<equiv> fmeasurable lebesgue"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1303
67982
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
  1304
lemma not_measurable_UNIV [simp]: "UNIV \<notin> lmeasurable"
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
  1305
  by (simp add: fmeasurable_def)
7643b005b29a various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
  1306
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1307
lemma\<^marker>\<open>tag important\<close> lmeasurable_iff_integrable:
63958
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1308
  "S \<in> lmeasurable \<longleftrightarrow> integrable lebesgue (indicator S :: 'a::euclidean_space \<Rightarrow> real)"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1309
  by (auto simp: fmeasurable_def integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1310
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1311
lemma lmeasurable_cbox [iff]: "cbox a b \<in> lmeasurable"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1312
  and lmeasurable_box [iff]: "box a b \<in> lmeasurable"
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1313
  by (auto simp: fmeasurable_def emeasure_lborel_box_eq emeasure_lborel_cbox_eq)
02de4a58e210 HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
hoelzl
parents: 63918
diff changeset
  1314
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
  1315
lemma
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
  1316
  fixes a::real
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
  1317
  shows lmeasurable_interval [iff]: "{a..b} \<in> lmeasurable" "{a<..<b} \<in> lmeasurable"
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
  1318
  apply (metis box_real(2) lmeasurable_cbox)
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
  1319
  by (metis box_real(1) lmeasurable_box)
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70381
diff changeset
  1320
67989
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1321
lemma fmeasurable_compact: "compact S \<Longrightarrow> S \<in> fmeasurable lborel"
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1322
  using emeasure_compact_finite[of S] by (intro fmeasurableI) (auto simp: borel_compact)
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1323
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1324
lemma lmeasurable_compact: "compact S \<Longrightarrow> S \<in> lmeasurable"
67989
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1325
  using fmeasurable_compact by (force simp: fmeasurable_def)
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1326
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1327
lemma measure_frontier:
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1328
   "bounded S \<Longrightarrow> measure lebesgue (frontier S) = measure lebesgue (closure S) - measure lebesgue (interior S)"
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1329
  using closure_subset interior_subset
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1330
  by (auto simp: frontier_def fmeasurable_compact intro!: measurable_measure_Diff)
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1331
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1332
lemma lmeasurable_closure:
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1333
   "bounded S \<Longrightarrow> closure S \<in> lmeasurable"
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1334
  by (simp add: lmeasurable_compact)
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1335
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1336
lemma lmeasurable_frontier:
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1337
   "bounded S \<Longrightarrow> frontier S \<in> lmeasurable"
706f86afff43 more results about measure and negligibility
paulson <lp15@cam.ac.uk>
parents: 67986
diff changeset
  1338
  by (simp add: compact_frontier_bounded lmeasurable_compact)
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1339
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1340
lemma lmeasurable_open: "bounded S \<Longrightarrow> open S \<Longrightarrow> S \<in> lmeasurable"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1341
  using emeasure_bounded_finite[of S] by (intro fmeasurableI) (auto simp: borel_open)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1342
67990
c0ebecf6e3eb some more random results
paulson <lp15@cam.ac.uk>
parents: 67989
diff changeset
  1343
lemma lmeasurable_ball [simp]: "ball a r \<in> lmeasurable"
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1344
  by (simp add: lmeasurable_open)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1345
67990
c0ebecf6e3eb some more random results
paulson <lp15@cam.ac.uk>
parents: 67989
diff changeset
  1346
lemma lmeasurable_cball [simp]: "cball a r \<in> lmeasurable"
c0ebecf6e3eb some more random results
paulson <lp15@cam.ac.uk>
parents: 67989
diff changeset
  1347
  by (simp add: lmeasurable_compact)
c0ebecf6e3eb some more random results
paulson <lp15@cam.ac.uk>
parents: 67989
diff changeset
  1348
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1349
lemma lmeasurable_interior: "bounded S \<Longrightarrow> interior S \<in> lmeasurable"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1350
  by (simp add: bounded_interior lmeasurable_open)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1351
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1352
lemma null_sets_cbox_Diff_box: "cbox a b - box a b \<in> null_sets lborel"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1353
proof -
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1354
  have "emeasure lborel (cbox a b - box a b) = 0"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1355
    by (subst emeasure_Diff) (auto simp: emeasure_lborel_cbox_eq emeasure_lborel_box_eq box_subset_cbox)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1356
  then have "cbox a b - box a b \<in> null_sets lborel"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1357
    by (auto simp: null_sets_def)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1358
  then show ?thesis
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1359
    by (auto dest!: AE_not_in)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1360
qed
67968
a5ad4c015d1c removed dots at the end of (sub)titles
nipkow
parents: 67673
diff changeset
  1361
67984
adc1a992c470 a few more results
paulson <lp15@cam.ac.uk>
parents: 67982
diff changeset
  1362
lemma bounded_set_imp_lmeasurable:
adc1a992c470 a few more results
paulson <lp15@cam.ac.uk>
parents: 67982
diff changeset
  1363
  assumes "bounded S" "S \<in> sets lebesgue" shows "S \<in> lmeasurable"
adc1a992c470 a few more results
paulson <lp15@cam.ac.uk>
parents: 67982
diff changeset
  1364
  by (metis assms bounded_Un emeasure_bounded_finite emeasure_completion fmeasurableI main_part_null_part_Un)
adc1a992c470 a few more results
paulson <lp15@cam.ac.uk>
parents: 67982
diff changeset
  1365
70381
b151d1f00204 More results about measure and integration theory
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  1366
lemma finite_measure_lebesgue_on: "S \<in> lmeasurable \<Longrightarrow> finite_measure (lebesgue_on S)"
b151d1f00204 More results about measure and integration theory
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  1367
  by (auto simp: finite_measureI fmeasurable_def emeasure_restrict_space)
b151d1f00204 More results about measure and integration theory
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  1368
b151d1f00204 More results about measure and integration theory
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  1369
lemma integrable_const_ivl [iff]:
b151d1f00204 More results about measure and integration theory
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  1370
  fixes a::"'a::ordered_euclidean_space"
b151d1f00204 More results about measure and integration theory
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  1371
  shows "integrable (lebesgue_on {a..b}) (\<lambda>x. c)"
b151d1f00204 More results about measure and integration theory
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  1372
  by (metis cbox_interval finite_measure.integrable_const finite_measure_lebesgue_on lmeasurable_cbox)
67986
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1373
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1374
subsection\<^marker>\<open>tag unimportant\<close>\<open>Translation preserves Lebesgue measure\<close>
67986
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1375
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1376
lemma sigma_sets_image:
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1377
  assumes S: "S \<in> sigma_sets \<Omega> M" and "M \<subseteq> Pow \<Omega>" "f ` \<Omega> = \<Omega>" "inj_on f \<Omega>"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1378
    and M: "\<And>y. y \<in> M \<Longrightarrow> f ` y \<in> M"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1379
  shows "(f ` S) \<in> sigma_sets \<Omega> M"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1380
  using S
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1381
proof (induct S rule: sigma_sets.induct)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1382
  case (Basic a) then show ?case
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1383
    by (simp add: M)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1384
next
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1385
  case Empty then show ?case
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1386
    by (simp add: sigma_sets.Empty)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1387
next
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1388
  case (Compl a)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1389
  then have "\<Omega> - a \<subseteq> \<Omega>" "a \<subseteq> \<Omega>"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1390
    by (auto simp: sigma_sets_into_sp [OF \<open>M \<subseteq> Pow \<Omega>\<close>])
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1391
  then show ?case
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1392
    by (auto simp: inj_on_image_set_diff [OF \<open>inj_on f \<Omega>\<close>] assms intro: Compl sigma_sets.Compl)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1393
next
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1394
  case (Union a) then show ?case
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1395
    by (metis image_UN sigma_sets.simps)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1396
qed
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1397
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1398
lemma null_sets_translation:
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1399
  assumes "N \<in> null_sets lborel" shows "{x. x - a \<in> N} \<in> null_sets lborel"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1400
proof -
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1401
  have [simp]: "(\<lambda>x. x + a) ` N = {x. x - a \<in> N}"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1402
    by force
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1403
  show ?thesis
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1404
    using assms emeasure_lebesgue_affine [of 1 a N] by (auto simp: null_sets_def)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1405
qed
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1406
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1407
lemma lebesgue_sets_translation:
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1408
  fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1409
  assumes S: "S \<in> sets lebesgue"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1410
  shows "((\<lambda>x. a + x) ` S) \<in> sets lebesgue"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1411
proof -
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1412
  have im_eq: "(+) a ` A = {x. x - a \<in> A}" for A
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1413
    by force
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1414
  have "((\<lambda>x. a + x) ` S) = ((\<lambda>x. -a + x) -` S) \<inter> (space lebesgue)"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1415
    using image_iff by fastforce
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1416
  also have "\<dots> \<in> sets lebesgue"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1417
  proof (rule measurable_sets [OF measurableI assms])
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1418
    fix A :: "'b set"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1419
    assume A: "A \<in> sets lebesgue"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1420
    have vim_eq: "(\<lambda>x. x - a) -` A = (+) a ` A" for A
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1421
      by force
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1422
    have "\<exists>s n N'. (+) a ` (S \<union> N) = s \<union> n \<and> s \<in> sets borel \<and> N' \<in> null_sets lborel \<and> n \<subseteq> N'"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1423
      if "S \<in> sets borel" and "N' \<in> null_sets lborel" and "N \<subseteq> N'" for S N N'
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1424
    proof (intro exI conjI)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1425
      show "(+) a ` (S \<union> N) = (\<lambda>x. a + x) ` S \<union> (\<lambda>x. a + x) ` N"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1426
        by auto
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1427
      show "(\<lambda>x. a + x) ` N' \<in> null_sets lborel"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1428
        using that by (auto simp: null_sets_translation im_eq)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1429
    qed (use that im_eq in auto)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1430
    with A have "(\<lambda>x. x - a) -` A \<in> sets lebesgue"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1431
      by (force simp: vim_eq completion_def intro!: sigma_sets_image)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1432
    then show "(+) (- a) -` A \<inter> space lebesgue \<in> sets lebesgue"
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1433
      by (auto simp: vimage_def im_eq)
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1434
  qed auto
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1435
  finally show ?thesis .
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1436
qed
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1437
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1438
lemma measurable_translation:
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1439
   "S \<in> lmeasurable \<Longrightarrow> ((+) a ` S) \<in> lmeasurable"
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1440
  using emeasure_lebesgue_affine [of 1 a S]
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1441
  apply (auto intro: lebesgue_sets_translation simp add: fmeasurable_def cong: image_cong_simp)
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1442
  apply (simp add: ac_simps)
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1443
  done
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1444
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1445
lemma measurable_translation_subtract:
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1446
   "S \<in> lmeasurable \<Longrightarrow> ((\<lambda>x. x - a) ` S) \<in> lmeasurable"
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1447
  using measurable_translation [of S "- a"] by (simp cong: image_cong_simp)
67986
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1448
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1449
lemma measure_translation:
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1450
  "measure lebesgue ((+) a ` S) = measure lebesgue S"
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1451
  using measure_lebesgue_affine [of 1 a S] by (simp add: ac_simps cong: image_cong_simp)
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1452
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1453
lemma measure_translation_subtract:
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1454
  "measure lebesgue ((\<lambda>x. x - a) ` S) = measure lebesgue S"
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1455
  using measure_translation [of "- a"] by (simp cong: image_cong_simp)
a03a63b81f44 tuned proofs
haftmann
parents: 69546
diff changeset
  1456
67986
b65c4a6a015e quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents: 67984
diff changeset
  1457
67968
a5ad4c015d1c removed dots at the end of (sub)titles
nipkow
parents: 67673
diff changeset
  1458
subsection \<open>A nice lemma for negligibility proofs\<close>
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1459
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1460
lemma summable_iff_suminf_neq_top: "(\<And>n. f n \<ge> 0) \<Longrightarrow> \<not> summable f \<Longrightarrow> (\<Sum>i. ennreal (f i)) = top"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1461
  by (metis summable_suminf_not_top)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1462
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1463
proposition\<^marker>\<open>tag important\<close> starlike_negligible_bounded_gmeasurable:
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1464
  fixes S :: "'a :: euclidean_space set"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1465
  assumes S: "S \<in> sets lebesgue" and "bounded S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1466
      and eq1: "\<And>c x. \<lbrakk>(c *\<^sub>R x) \<in> S; 0 \<le> c; x \<in> S\<rbrakk> \<Longrightarrow> c = 1"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1467
    shows "S \<in> null_sets lebesgue"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1468
proof -
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1469
  obtain M where "0 < M" "S \<subseteq> ball 0 M"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1470
    using \<open>bounded S\<close> by (auto dest: bounded_subset_ballD)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1471
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1472
  let ?f = "\<lambda>n. root DIM('a) (Suc n)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1473
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1474
  have vimage_eq_image: "(*\<^sub>R) (?f n) -` S = (*\<^sub>R) (1 / ?f n) ` S" for n
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1475
    apply safe
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1476
    subgoal for x by (rule image_eqI[of _ _ "?f n *\<^sub>R x"]) auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1477
    subgoal by auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1478
    done
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1479
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1480
  have eq: "(1 / ?f n) ^ DIM('a) = 1 / Suc n" for n
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1481
    by (simp add: field_simps)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1482
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1483
  { fix n x assume x: "root DIM('a) (1 + real n) *\<^sub>R x \<in> S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1484
    have "1 * norm x \<le> root DIM('a) (1 + real n) * norm x"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1485
      by (rule mult_mono) auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1486
    also have "\<dots> < M"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1487
      using x \<open>S \<subseteq> ball 0 M\<close> by auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1488
    finally have "norm x < M" by simp }
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1489
  note less_M = this
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1490
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1491
  have "(\<Sum>n. ennreal (1 / Suc n)) = top"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1492
    using not_summable_harmonic[where 'a=real] summable_Suc_iff[where f="\<lambda>n. 1 / (real n)"]
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1493
    by (intro summable_iff_suminf_neq_top) (auto simp add: inverse_eq_divide)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1494
  then have "top * emeasure lebesgue S = (\<Sum>n. (1 / ?f n)^DIM('a) * emeasure lebesgue S)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1495
    unfolding ennreal_suminf_multc eq by simp
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1496
  also have "\<dots> = (\<Sum>n. emeasure lebesgue ((*\<^sub>R) (?f n) -` S))"
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1497
    unfolding vimage_eq_image using emeasure_lebesgue_affine[of "1 / ?f n" 0 S for n] by simp
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1498
  also have "\<dots> = emeasure lebesgue (\<Union>n. (*\<^sub>R) (?f n) -` S)"
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1499
  proof (intro suminf_emeasure)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1500
    show "disjoint_family (\<lambda>n. (*\<^sub>R) (?f n) -` S)"
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1501
      unfolding disjoint_family_on_def
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1502
    proof safe
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1503
      fix m n :: nat and x assume "m \<noteq> n" "?f m *\<^sub>R x \<in> S" "?f n *\<^sub>R x \<in> S"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1504
      with eq1[of "?f m / ?f n" "?f n *\<^sub>R x"] show "x \<in> {}"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1505
        by auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1506
    qed
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1507
    have "(*\<^sub>R) (?f i) -` S \<in> sets lebesgue" for i
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1508
      using measurable_sets[OF lebesgue_measurable_scaling[of "?f i"] S] by auto
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68403
diff changeset
  1509
    then show "range (\<lambda>i. (*\<^sub>R) (?f i) -` S) \<subseteq> sets lebesgue"
63959
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1510
      by auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1511
  qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1512
  also have "\<dots> \<le> emeasure lebesgue (ball 0 M :: 'a set)"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1513
    using less_M by (intro emeasure_mono) auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1514
  also have "\<dots> < top"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1515
    using lmeasurable_ball by (auto simp: fmeasurable_def)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1516
  finally have "emeasure lebesgue S = 0"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1517
    by (simp add: ennreal_top_mult split: if_split_asm)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1518
  then show "S \<in> null_sets lebesgue"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1519
    unfolding null_sets_def using \<open>S \<in> sets lebesgue\<close> by auto
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1520
qed
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1521
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1522
corollary starlike_negligible_compact:
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1523
  "compact S \<Longrightarrow> (\<And>c x. \<lbrakk>(c *\<^sub>R x) \<in> S; 0 \<le> c; x \<in> S\<rbrakk> \<Longrightarrow> c = 1) \<Longrightarrow> S \<in> null_sets lebesgue"
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1524
  using starlike_negligible_bounded_gmeasurable[of S] by (auto simp: compact_eq_bounded_closed)
f77dca1abf1b HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63958
diff changeset
  1525
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1526
proposition outer_regular_lborel_le:
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1527
  assumes B[measurable]: "B \<in> sets borel" and "0 < (e::real)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1528
  obtains U where "open U" "B \<subseteq> U" and "emeasure lborel (U - B) \<le> e"
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1529
proof -
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1530
  let ?\<mu> = "emeasure lborel"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1531
  let ?B = "\<lambda>n::nat. ball 0 n :: 'a set"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1532
  let ?e = "\<lambda>n. e*((1/2)^Suc n)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1533
  have "\<forall>n. \<exists>U. open U \<and> ?B n \<inter> B \<subseteq> U \<and> ?\<mu> (U - B) < ?e n"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1534
  proof
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1535
    fix n :: nat
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1536
    let ?A = "density lborel (indicator (?B n))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1537
    have emeasure_A: "X \<in> sets borel \<Longrightarrow> emeasure ?A X = ?\<mu> (?B n \<inter> X)" for X
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1538
      by (auto simp: emeasure_density borel_measurable_indicator indicator_inter_arith[symmetric])
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1539
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1540
    have finite_A: "emeasure ?A (space ?A) \<noteq> \<infinity>"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1541
      using emeasure_bounded_finite[of "?B n"] by (auto simp: emeasure_A)
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1542
    interpret A: finite_measure ?A
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1543
      by rule fact
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69064
diff changeset
  1544
    have "emeasure ?A B + ?e n > (INF U\<in>{U. B \<subseteq> U \<and> open U}. emeasure ?A U)"
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1545
      using \<open>0<e\<close> by (auto simp: outer_regular[OF _ finite_A B, symmetric])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1546
    then obtain U where U: "B \<subseteq> U" "open U" and muU: "?\<mu> (?B n \<inter> B) + ?e n > ?\<mu> (?B n \<inter> U)"
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1547
      unfolding INF_less_iff by (auto simp: emeasure_A)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1548
    moreover
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1549
    { have "?\<mu> ((?B n \<inter> U) - B) = ?\<mu> ((?B n \<inter> U) - (?B n \<inter> B))"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1550
        using U by (intro arg_cong[where f="?\<mu>"]) auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1551
      also have "\<dots> = ?\<mu> (?B n \<inter> U) - ?\<mu> (?B n \<inter> B)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1552
        using U A.emeasure_finite[of B]
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1553
        by (intro emeasure_Diff) (auto simp del: A.emeasure_finite simp: emeasure_A)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1554
      also have "\<dots> < ?e n"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1555
        using U muU A.emeasure_finite[of B]
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1556
        by (subst minus_less_iff_ennreal)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1557
          (auto simp del: A.emeasure_finite simp: emeasure_A less_top ac_simps intro!: emeasure_mono)
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1558
      finally have "?\<mu> ((?B n \<inter> U) - B) < ?e n" . }
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1559
    ultimately show "\<exists>U. open U \<and> ?B n \<inter> B \<subseteq> U \<and> ?\<mu> (U - B) < ?e n"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1560
      by (intro exI[of _ "?B n \<inter> U"]) auto
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1561
  qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1562
  then obtain U
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1563
    where U: "\<And>n. open (U n)" "\<And>n. ?B n \<inter> B \<subseteq> U n" "\<And>n. ?\<mu> (U n - B) < ?e n"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1564
    by metis
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1565
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1566
  proof
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1567
    { fix x assume "x \<in> B"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1568
      moreover
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1569
      obtain n where "norm x < real n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1570
        using reals_Archimedean2 by blast
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1571
      ultimately have "x \<in> (\<Union>n. U n)"
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1572
        using U(2)[of n] by auto }
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1573
    note * = this
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1574
    then show "open (\<Union>n. U n)" "B \<subseteq> (\<Union>n. U n)"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1575
      using U by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1576
    have "?\<mu> (\<Union>n. U n - B) \<le> (\<Sum>n. ?\<mu> (U n - B))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1577
      using U(1) by (intro emeasure_subadditive_countably) auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1578
    also have "\<dots> \<le> (\<Sum>n. ennreal (?e n))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1579
      using U(3) by (intro suminf_le) (auto intro: less_imp_le)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1580
    also have "\<dots> = ennreal (e * 1)"
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1581
      using \<open>0<e\<close> by (intro suminf_ennreal_eq sums_mult power_half_series) auto
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1582
    finally show "emeasure lborel ((\<Union>n. U n) - B) \<le> ennreal e"
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1583
      by simp
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1584
  qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1585
qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1586
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1587
lemma\<^marker>\<open>tag important\<close> outer_regular_lborel:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1588
  assumes B: "B \<in> sets borel" and "0 < (e::real)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1589
  obtains U where "open U" "B \<subseteq> U" "emeasure lborel (U - B) < e"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1590
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1591
  obtain U where U: "open U" "B \<subseteq> U" and "emeasure lborel (U-B) \<le> e/2"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1592
    using outer_regular_lborel_le [OF B, of "e/2"] \<open>e > 0\<close>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1593
    by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1594
  moreover have "ennreal (e/2) < ennreal e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1595
    using \<open>e > 0\<close> by (simp add: ennreal_lessI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1596
  ultimately have "emeasure lborel (U-B) < e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1597
    by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1598
  with U show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1599
    using that by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1600
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1601
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1602
lemma completion_upper:
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1603
  assumes A: "A \<in> sets (completion M)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1604
  obtains A' where "A \<subseteq> A'" "A' \<in> sets M" "A' - A \<in> null_sets (completion M)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1605
                   "emeasure (completion M) A = emeasure M A'"
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1606
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1607
  from AE_notin_null_part[OF A] obtain N where N: "N \<in> null_sets M" "null_part M A \<subseteq> N"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1608
    unfolding eventually_ae_filter using null_part_null_sets[OF A, THEN null_setsD2, THEN sets.sets_into_space] by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1609
  let ?A' = "main_part M A \<union> N"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1610
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1611
  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1612
    show "A \<subseteq> ?A'"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1613
      using \<open>null_part M A \<subseteq> N\<close> by (subst main_part_null_part_Un[symmetric, OF A]) auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1614
    have "main_part M A \<subseteq> A"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1615
      using assms main_part_null_part_Un by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1616
    then have "?A' - A \<subseteq> N"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1617
      by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1618
    with N show "?A' - A \<in> null_sets (completion M)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1619
      by (blast intro: null_sets_completionI completion.complete_measure_axioms complete_measure.complete2)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1620
    show "emeasure (completion M) A = emeasure M (main_part M A \<union> N)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1621
      using A \<open>N \<in> null_sets M\<close> by (simp add: emeasure_Un_null_set)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1622
  qed (use A N in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1623
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1624
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1625
lemma sets_lebesgue_outer_open:
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1626
  fixes e::real
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1627
  assumes S: "S \<in> sets lebesgue" and "e > 0"
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1628
  obtains T where "open T" "S \<subseteq> T" "(T - S) \<in> lmeasurable" "emeasure lebesgue (T - S) < ennreal e"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1629
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1630
  obtain S' where S': "S \<subseteq> S'" "S' \<in> sets borel"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1631
              and null: "S' - S \<in> null_sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1632
              and em: "emeasure lebesgue S = emeasure lborel S'"
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1633
    using completion_upper[of S lborel] S by auto
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1634
  then have f_S': "S' \<in> sets borel"
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1635
    using S by (auto simp: fmeasurable_def)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1636
  with outer_regular_lborel[OF _ \<open>0<e\<close>]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1637
  obtain U where U: "open U" "S' \<subseteq> U" "emeasure lborel (U - S') < e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1638
    by blast
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1639
  show thesis
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1640
  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1641
    show "open U" "S \<subseteq> U"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1642
      using f_S' U S' by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1643
  have "(U - S) = (U - S') \<union> (S' - S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1644
    using S' U by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1645
  then have eq: "emeasure lebesgue (U - S) = emeasure lborel (U - S')"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1646
    using null  by (simp add: U(1) emeasure_Un_null_set f_S' sets.Diff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1647
  have "(U - S) \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1648
    by (simp add: S U(1) sets.Diff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1649
  then show "(U - S) \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents: 67990
diff changeset
  1650
    unfolding fmeasurable_def using U(3) eq less_le_trans by fastforce
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1651
  with eq U show "emeasure lebesgue (U - S) < e"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1652
    by (simp add: eq)
63968
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1653
  qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1654
qed
4359400adfe7 HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
hoelzl
parents: 63959
diff changeset
  1655
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1656
lemma sets_lebesgue_inner_closed:
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1657
  fixes e::real
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1658
  assumes "S \<in> sets lebesgue" "e > 0"
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1659
  obtains T where "closed T" "T \<subseteq> S" "S-T \<in> lmeasurable" "emeasure lebesgue (S - T) < ennreal e"
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1660
proof -
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1661
  have "-S \<in> sets lebesgue"
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1662
    using assms by (simp add: Compl_in_sets_lebesgue)
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1663
  then obtain T where "open T" "-S \<subseteq> T"
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1664
          and T: "(T - -S) \<in> lmeasurable" "emeasure lebesgue (T - -S) < e"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1665
    using sets_lebesgue_outer_open assms  by blast
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1666
  show thesis
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1667
  proof
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1668
    show "closed (-T)"
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1669
      using \<open>open T\<close> by blast
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1670
    show "-T \<subseteq> S"
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1671
      using \<open>- S \<subseteq> T\<close> by auto
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1672
    show "S - ( -T) \<in> lmeasurable" "emeasure lebesgue (S - (- T)) < e"
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1673
      using T by (auto simp: Int_commute)
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1674
  qed
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1675
qed
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
  1676
67673
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
  1677
lemma lebesgue_openin:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1678
   "\<lbrakk>openin (top_of_set S) T; S \<in> sets lebesgue\<rbrakk> \<Longrightarrow> T \<in> sets lebesgue"
67673
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
  1679
  by (metis borel_open openin_open sets.Int sets_completionI_sets sets_lborel)
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
  1680
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
  1681
lemma lebesgue_closedin:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1682
   "\<lbrakk>closedin (top_of_set S) T; S \<in> sets lebesgue\<rbrakk> \<Longrightarrow> T \<in> sets lebesgue"
67673
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
  1683
  by (metis borel_closed closedin_closed sets.Int sets_completionI_sets sets_lborel)
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
  1684
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1685
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1686
subsection\<open>\<open>F_sigma\<close> and \<open>G_delta\<close> sets.\<close> 
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1687
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1688
\<comment> \<open>\<^url>\<open>https://en.wikipedia.org/wiki/F-sigma_set\<close>\<close>
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1689
inductive\<^marker>\<open>tag important\<close> fsigma :: "'a::topological_space set \<Rightarrow> bool" where
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1690
  "(\<And>n::nat. closed (F n)) \<Longrightarrow> fsigma (\<Union>(F ` UNIV))"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1691
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1692
inductive\<^marker>\<open>tag important\<close> gdelta :: "'a::topological_space set \<Rightarrow> bool" where
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1693
  "(\<And>n::nat. open (F n)) \<Longrightarrow> gdelta (\<Inter>(F ` UNIV))"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1694
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1695
lemma fsigma_Union_compact:
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1696
  fixes S :: "'a::{real_normed_vector,heine_borel} set"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1697
  shows "fsigma S \<longleftrightarrow> (\<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> Collect compact \<and> S = \<Union>(F ` UNIV))"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1698
proof safe
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1699
  assume "fsigma S"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1700
  then obtain F :: "nat \<Rightarrow> 'a set" where F: "range F \<subseteq> Collect closed" "S = \<Union>(F ` UNIV)"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1701
    by (meson fsigma.cases image_subsetI mem_Collect_eq)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1702
  then have "\<exists>D::nat \<Rightarrow> 'a set. range D \<subseteq> Collect compact \<and> \<Union>(D ` UNIV) = F i" for i
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1703
    using closed_Union_compact_subsets [of "F i"]
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1704
    by (metis image_subsetI mem_Collect_eq range_subsetD)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1705
  then obtain D :: "nat \<Rightarrow> nat \<Rightarrow> 'a set"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1706
    where D: "\<And>i. range (D i) \<subseteq> Collect compact \<and> \<Union>((D i) ` UNIV) = F i"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1707
    by metis
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1708
  let ?DD = "\<lambda>n. (\<lambda>(i,j). D i j) (prod_decode n)"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1709
  show "\<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> Collect compact \<and> S = \<Union>(F ` UNIV)"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1710
  proof (intro exI conjI)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1711
    show "range ?DD \<subseteq> Collect compact"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1712
      using D by clarsimp (metis mem_Collect_eq rangeI split_conv subsetCE surj_pair)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1713
    show "S = \<Union> (range ?DD)"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1714
    proof
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1715
      show "S \<subseteq> \<Union> (range ?DD)"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1716
        using D F
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1717
        by clarsimp (metis UN_iff old.prod.case prod_decode_inverse prod_encode_eq)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1718
      show "\<Union> (range ?DD) \<subseteq> S"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1719
        using D F  by fastforce
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1720
    qed
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1721
  qed
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1722
next
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1723
  fix F :: "nat \<Rightarrow> 'a set"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1724
  assume "range F \<subseteq> Collect compact" and "S = \<Union>(F ` UNIV)"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1725
  then show "fsigma (\<Union>(F ` UNIV))"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1726
    by (simp add: compact_imp_closed fsigma.intros image_subset_iff)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1727
qed
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1728
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1729
lemma gdelta_imp_fsigma: "gdelta S \<Longrightarrow> fsigma (- S)"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1730
proof (induction rule: gdelta.induct)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1731
  case (1 F)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1732
  have "- \<Inter>(F ` UNIV) = (\<Union>i. -(F i))"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1733
    by auto
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1734
  then show ?case
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1735
    by (simp add: fsigma.intros closed_Compl 1)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1736
qed
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1737
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1738
lemma fsigma_imp_gdelta: "fsigma S \<Longrightarrow> gdelta (- S)"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1739
proof (induction rule: fsigma.induct)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1740
  case (1 F)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1741
  have "- \<Union>(F ` UNIV) = (\<Inter>i. -(F i))"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1742
    by auto
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1743
  then show ?case
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1744
    by (simp add: 1 gdelta.intros open_closed)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1745
qed
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1746
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1747
lemma gdelta_complement: "gdelta(- S) \<longleftrightarrow> fsigma S"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1748
  using fsigma_imp_gdelta gdelta_imp_fsigma by force
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70271
diff changeset
  1749
70380
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1750
lemma lebesgue_set_almost_fsigma:
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1751
  assumes "S \<in> sets lebesgue"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1752
  obtains C T where "fsigma C" "T \<in> null_sets lebesgue" "C \<union> T = S" "disjnt C T"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1753
proof -
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1754
  { fix n::nat
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1755
    obtain T where "closed T" "T \<subseteq> S" "S-T \<in> lmeasurable" "emeasure lebesgue (S - T) < ennreal (1 / Suc n)"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1756
      using sets_lebesgue_inner_closed [OF assms]
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1757
      by (metis of_nat_0_less_iff zero_less_Suc zero_less_divide_1_iff)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1758
    then have "\<exists>T. closed T \<and> T \<subseteq> S \<and> S - T \<in> lmeasurable \<and> measure lebesgue (S-T) < 1 / Suc n"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1759
      by (metis emeasure_eq_measure2 ennreal_leI not_le)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1760
  }
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1761
  then obtain F where F: "\<And>n::nat. closed (F n) \<and> F n \<subseteq> S \<and> S - F n \<in> lmeasurable \<and> measure lebesgue (S - F n) < 1 / Suc n"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1762
    by metis
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1763
  let ?C = "\<Union>(F ` UNIV)"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1764
  show thesis
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1765
  proof
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1766
    show "fsigma ?C"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1767
      using F by (simp add: fsigma.intros)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1768
    show "(S - ?C) \<in> null_sets lebesgue"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1769
    proof (clarsimp simp add: completion.null_sets_outer_le)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1770
      fix e :: "real"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1771
      assume "0 < e"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1772
      then obtain n where n: "1 / Suc n < e"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1773
        using nat_approx_posE by metis
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1774
      show "\<exists>T \<in> lmeasurable. S - (\<Union>x. F x) \<subseteq> T \<and> measure lebesgue T \<le> e"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1775
      proof (intro bexI conjI)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1776
        show "measure lebesgue (S - F n) \<le> e"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1777
          by (meson F n less_trans not_le order.asym)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1778
      qed (use F in auto)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1779
    qed
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1780
    show "?C \<union> (S - ?C) = S"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1781
      using F by blast
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1782
    show "disjnt ?C (S - ?C)"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1783
      by (auto simp: disjnt_def)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1784
  qed
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1785
qed
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1786
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1787
lemma lebesgue_set_almost_gdelta:
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1788
  assumes "S \<in> sets lebesgue"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1789
  obtains C T where "gdelta C" "T \<in> null_sets lebesgue" "S \<union> T = C" "disjnt S T"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1790
proof -
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1791
  have "-S \<in> sets lebesgue"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1792
    using assms Compl_in_sets_lebesgue by blast
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1793
  then obtain C T where C: "fsigma C" "T \<in> null_sets lebesgue" "C \<union> T = -S" "disjnt C T"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1794
    using lebesgue_set_almost_fsigma by metis
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1795
  show thesis
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1796
  proof
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1797
    show "gdelta (-C)"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1798
      by (simp add: \<open>fsigma C\<close> fsigma_imp_gdelta)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1799
    show "S \<union> T = -C" "disjnt S T"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1800
      using C by (auto simp: disjnt_def)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1801
  qed (use C in auto)
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1802
qed
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  1803
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1804
end