src/HOL/Probability/Borel_Space.thy
author hoelzl
Wed, 10 Oct 2012 12:12:34 +0200
changeset 49799 15ea98537c76
parent 49774 dfa8ddb874ce
child 50001 382bd3173584
permissions -rw-r--r--
strong nonnegativ (instead of ae nn) for induction rule
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42150
b0c0638c4aad tuned headers;
wenzelm
parents: 42067
diff changeset
     1
(*  Title:      HOL/Probability/Borel_Space.thy
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     2
    Author:     Johannes Hölzl, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     3
    Author:     Armin Heller, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     4
*)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
     5
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
     6
header {*Borel spaces*}
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
     7
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
     8
theory Borel_Space
45288
fc3c7db5bb2f correct import path
hoelzl
parents: 45287
diff changeset
     9
  imports Sigma_Algebra "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    10
begin
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    11
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    12
section "Generic Borel spaces"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    13
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    14
definition borel :: "'a::topological_space measure" where
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    15
  "borel = sigma UNIV {S. open S}"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    16
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    17
abbreviation "borel_measurable M \<equiv> measurable M borel"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    18
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    19
lemma in_borel_measurable:
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    20
   "f \<in> borel_measurable M \<longleftrightarrow>
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    21
    (\<forall>S \<in> sigma_sets UNIV {S. open S}. f -` S \<inter> space M \<in> sets M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    22
  by (auto simp add: measurable_def borel_def)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    23
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    24
lemma in_borel_measurable_borel:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    25
   "f \<in> borel_measurable M \<longleftrightarrow>
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    26
    (\<forall>S \<in> sets borel.
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    27
      f -` S \<inter> space M \<in> sets M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    28
  by (auto simp add: measurable_def borel_def)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    29
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    30
lemma space_borel[simp]: "space borel = UNIV"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    31
  unfolding borel_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    32
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    33
lemma borel_open[simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    34
  assumes "open A" shows "A \<in> sets borel"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    35
proof -
44537
c10485a6a7af make HOL-Probability respect set/pred distinction
huffman
parents: 44282
diff changeset
    36
  have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms .
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    37
  thus ?thesis unfolding borel_def by auto
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    38
qed
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    39
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    40
lemma borel_closed[simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    41
  assumes "closed A" shows "A \<in> sets borel"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    42
proof -
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    43
  have "space borel - (- A) \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    44
    using assms unfolding closed_def by (blast intro: borel_open)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    45
  thus ?thesis by simp
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    46
qed
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    47
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
    48
lemma borel_comp[intro,simp]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    49
  unfolding Compl_eq_Diff_UNIV by (intro Diff) auto
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
    50
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    51
lemma borel_measurable_vimage:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    52
  fixes f :: "'a \<Rightarrow> 'x::t2_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    53
  assumes borel: "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    54
  shows "f -` {x} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    55
proof (cases "x \<in> f ` space M")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    56
  case True then obtain y where "x = f y" by auto
41969
1cf3e4107a2a moved t2_spaces to HOL image
hoelzl
parents: 41830
diff changeset
    57
  from closed_singleton[of "f y"]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    58
  have "{f y} \<in> sets borel" by (rule borel_closed)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    59
  with assms show ?thesis
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    60
    unfolding in_borel_measurable_borel `x = f y` by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    61
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    62
  case False hence "f -` {x} \<inter> space M = {}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    63
  thus ?thesis by auto
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    64
qed
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    65
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    66
lemma borel_measurableI:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    67
  fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    68
  assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    69
  shows "f \<in> borel_measurable M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    70
  unfolding borel_def
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    71
proof (rule measurable_measure_of, simp_all)
44537
c10485a6a7af make HOL-Probability respect set/pred distinction
huffman
parents: 44282
diff changeset
    72
  fix S :: "'x set" assume "open S" thus "f -` S \<inter> space M \<in> sets M"
c10485a6a7af make HOL-Probability respect set/pred distinction
huffman
parents: 44282
diff changeset
    73
    using assms[of S] by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    74
qed
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    75
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    76
lemma borel_singleton[simp, intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    77
  fixes x :: "'a::t1_space"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    78
  shows "A \<in> sets borel \<Longrightarrow> insert x A \<in> sets borel"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    79
  proof (rule insert_in_sets)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    80
    show "{x} \<in> sets borel"
41969
1cf3e4107a2a moved t2_spaces to HOL image
hoelzl
parents: 41830
diff changeset
    81
      using closed_singleton[of x] by (rule borel_closed)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    82
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    83
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    84
lemma borel_measurable_const[simp, intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    85
  "(\<lambda>x. c) \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    86
  by auto
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    87
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    88
lemma borel_measurable_indicator[simp, intro!]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    89
  assumes A: "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    90
  shows "indicator A \<in> borel_measurable M"
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46884
diff changeset
    91
  unfolding indicator_def [abs_def] using A
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    92
  by (auto intro!: measurable_If_set)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    93
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    94
lemma borel_measurable_indicator_iff:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    95
  "(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    96
    (is "?I \<in> borel_measurable M \<longleftrightarrow> _")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    97
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    98
  assume "?I \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    99
  then have "?I -` {1} \<inter> space M \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   100
    unfolding measurable_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   101
  also have "?I -` {1} \<inter> space M = A \<inter> space M"
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46884
diff changeset
   102
    unfolding indicator_def [abs_def] by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   103
  finally show "A \<inter> space M \<in> sets M" .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   104
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   105
  assume "A \<inter> space M \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   106
  moreover have "?I \<in> borel_measurable M \<longleftrightarrow>
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   107
    (indicator (A \<inter> space M) :: 'a \<Rightarrow> 'x) \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   108
    by (intro measurable_cong) (auto simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   109
  ultimately show "?I \<in> borel_measurable M" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   110
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   111
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   112
lemma borel_measurable_subalgebra:
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41097
diff changeset
   113
  assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
   114
  shows "f \<in> borel_measurable M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
   115
  using assms unfolding measurable_def by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
   116
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   117
section "Borel spaces on euclidean spaces"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   118
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   119
lemma lessThan_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   120
  fixes a :: "'a\<Colon>ordered_euclidean_space"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   121
  shows "{..< a} \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   122
  by (blast intro: borel_open)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   123
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   124
lemma greaterThan_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   125
  fixes a :: "'a\<Colon>ordered_euclidean_space"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   126
  shows "{a <..} \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   127
  by (blast intro: borel_open)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   128
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   129
lemma greaterThanLessThan_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   130
  fixes a b :: "'a\<Colon>ordered_euclidean_space"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   131
  shows "{a<..<b} \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   132
  by (blast intro: borel_open)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   133
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   134
lemma atMost_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   135
  fixes a :: "'a\<Colon>ordered_euclidean_space"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   136
  shows "{..a} \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   137
  by (blast intro: borel_closed)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   138
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   139
lemma atLeast_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   140
  fixes a :: "'a\<Colon>ordered_euclidean_space"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   141
  shows "{a..} \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   142
  by (blast intro: borel_closed)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   143
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   144
lemma atLeastAtMost_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   145
  fixes a b :: "'a\<Colon>ordered_euclidean_space"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   146
  shows "{a..b} \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   147
  by (blast intro: borel_closed)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   148
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   149
lemma greaterThanAtMost_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   150
  fixes a b :: "'a\<Colon>ordered_euclidean_space"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   151
  shows "{a<..b} \<in> sets borel"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   152
  unfolding greaterThanAtMost_def by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   153
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   154
lemma atLeastLessThan_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   155
  fixes a b :: "'a\<Colon>ordered_euclidean_space"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   156
  shows "{a..<b} \<in> sets borel"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   157
  unfolding atLeastLessThan_def by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   158
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   159
lemma hafspace_less_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   160
  fixes a :: real
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   161
  shows "{x::'a::euclidean_space. a < x $$ i} \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   162
  by (auto intro!: borel_open open_halfspace_component_gt)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   163
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   164
lemma hafspace_greater_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   165
  fixes a :: real
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   166
  shows "{x::'a::euclidean_space. x $$ i < a} \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   167
  by (auto intro!: borel_open open_halfspace_component_lt)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   168
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   169
lemma hafspace_less_eq_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   170
  fixes a :: real
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   171
  shows "{x::'a::euclidean_space. a \<le> x $$ i} \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   172
  by (auto intro!: borel_closed closed_halfspace_component_ge)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   173
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   174
lemma hafspace_greater_eq_borel[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   175
  fixes a :: real
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   176
  shows "{x::'a::euclidean_space. x $$ i \<le> a} \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   177
  by (auto intro!: borel_closed closed_halfspace_component_le)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   178
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   179
lemma borel_measurable_less[simp, intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   180
  fixes f :: "'a \<Rightarrow> real"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   181
  assumes f: "f \<in> borel_measurable M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   182
  assumes g: "g \<in> borel_measurable M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   183
  shows "{w \<in> space M. f w < g w} \<in> sets M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   184
proof -
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   185
  have "{w \<in> space M. f w < g w} =
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   186
        (\<Union>r. (f -` {..< of_rat r} \<inter> space M) \<inter> (g -` {of_rat r <..} \<inter> space M))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   187
    using Rats_dense_in_real by (auto simp add: Rats_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   188
  then show ?thesis using f g
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   189
    by simp (blast intro: measurable_sets)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   190
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   191
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   192
lemma borel_measurable_le[simp, intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   193
  fixes f :: "'a \<Rightarrow> real"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   194
  assumes f: "f \<in> borel_measurable M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   195
  assumes g: "g \<in> borel_measurable M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   196
  shows "{w \<in> space M. f w \<le> g w} \<in> sets M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   197
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   198
  have "{w \<in> space M. f w \<le> g w} = space M - {w \<in> space M. g w < f w}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   199
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   200
  thus ?thesis using f g
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   201
    by simp blast
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   202
qed
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   203
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   204
lemma borel_measurable_eq[simp, intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   205
  fixes f :: "'a \<Rightarrow> real"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   206
  assumes f: "f \<in> borel_measurable M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   207
  assumes g: "g \<in> borel_measurable M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   208
  shows "{w \<in> space M. f w = g w} \<in> sets M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   209
proof -
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   210
  have "{w \<in> space M. f w = g w} =
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33535
diff changeset
   211
        {w \<in> space M. f w \<le> g w} \<inter> {w \<in> space M. g w \<le> f w}"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   212
    by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   213
  thus ?thesis using f g by auto
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   214
qed
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   215
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   216
lemma borel_measurable_neq[simp, intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   217
  fixes f :: "'a \<Rightarrow> real"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   218
  assumes f: "f \<in> borel_measurable M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   219
  assumes g: "g \<in> borel_measurable M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   220
  shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   221
proof -
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   222
  have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}"
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
   223
    by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   224
  thus ?thesis using f g by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   225
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   226
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   227
subsection "Borel space equals sigma algebras over intervals"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   228
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   229
lemma rational_boxes:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   230
  fixes x :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   231
  assumes "0 < e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   232
  shows "\<exists>a b. (\<forall>i. a $$ i \<in> \<rat>) \<and> (\<forall>i. b $$ i \<in> \<rat>) \<and> x \<in> {a <..< b} \<and> {a <..< b} \<subseteq> ball x e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   233
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   234
  def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   235
  then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   236
  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x $$ i \<and> x $$ i - y < e'" (is "\<forall>i. ?th i")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   237
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   238
    fix i from Rats_dense_in_real[of "x $$ i - e'" "x $$ i"] e
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   239
    show "?th i" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   240
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   241
  from choice[OF this] guess a .. note a = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   242
  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x $$ i < y \<and> y - x $$ i < e'" (is "\<forall>i. ?th i")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   243
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   244
    fix i from Rats_dense_in_real[of "x $$ i" "x $$ i + e'"] e
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   245
    show "?th i" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   246
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   247
  from choice[OF this] guess b .. note b = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   248
  { fix y :: 'a assume *: "Chi a < y" "y < Chi b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   249
    have "dist x y = sqrt (\<Sum>i<DIM('a). (dist (x $$ i) (y $$ i))\<twosuperior>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   250
      unfolding setL2_def[symmetric] by (rule euclidean_dist_l2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   251
    also have "\<dots> < sqrt (\<Sum>i<DIM('a). e^2 / real (DIM('a)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   252
    proof (rule real_sqrt_less_mono, rule setsum_strict_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   253
      fix i assume i: "i \<in> {..<DIM('a)}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   254
      have "a i < y$$i \<and> y$$i < b i" using * i eucl_less[where 'a='a] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   255
      moreover have "a i < x$$i" "x$$i - a i < e'" using a by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   256
      moreover have "x$$i < b i" "b i - x$$i < e'" using b by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   257
      ultimately have "\<bar>x$$i - y$$i\<bar> < 2 * e'" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   258
      then have "dist (x $$ i) (y $$ i) < e/sqrt (real (DIM('a)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   259
        unfolding e'_def by (auto simp: dist_real_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   260
      then have "(dist (x $$ i) (y $$ i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   261
        by (rule power_strict_mono) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   262
      then show "(dist (x $$ i) (y $$ i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   263
        by (simp add: power_divide)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   264
    qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   265
    also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat DIM_positive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   266
    finally have "dist x y < e" . }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   267
  with a b show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   268
    apply (rule_tac exI[of _ "Chi a"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   269
    apply (rule_tac exI[of _ "Chi b"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   270
    using eucl_less[where 'a='a] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   271
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   272
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   273
lemma ex_rat_list:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   274
  fixes x :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   275
  assumes "\<And> i. x $$ i \<in> \<rat>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   276
  shows "\<exists> r. length r = DIM('a) \<and> (\<forall> i < DIM('a). of_rat (r ! i) = x $$ i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   277
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   278
  have "\<forall>i. \<exists>r. x $$ i = of_rat r" using assms unfolding Rats_def by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   279
  from choice[OF this] guess r ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   280
  then show ?thesis by (auto intro!: exI[of _ "map r [0 ..< DIM('a)]"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   281
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   282
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   283
lemma open_UNION:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   284
  fixes M :: "'a\<Colon>ordered_euclidean_space set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   285
  assumes "open M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   286
  shows "M = UNION {(a, b) | a b. {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)} \<subseteq> M}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   287
                   (\<lambda> (a, b). {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   288
    (is "M = UNION ?idx ?box")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   289
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   290
  fix x assume "x \<in> M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   291
  obtain e where e: "e > 0" "ball x e \<subseteq> M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   292
    using openE[OF assms `x \<in> M`] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   293
  then obtain a b where ab: "x \<in> {a <..< b}" "\<And>i. a $$ i \<in> \<rat>" "\<And>i. b $$ i \<in> \<rat>" "{a <..< b} \<subseteq> ball x e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   294
    using rational_boxes[OF e(1)] by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   295
  then obtain p q where pq: "length p = DIM ('a)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   296
                            "length q = DIM ('a)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   297
                            "\<forall> i < DIM ('a). of_rat (p ! i) = a $$ i \<and> of_rat (q ! i) = b $$ i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   298
    using ex_rat_list[OF ab(2)] ex_rat_list[OF ab(3)] by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   299
  hence p: "Chi (of_rat \<circ> op ! p) = a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   300
    using euclidean_eq[of "Chi (of_rat \<circ> op ! p)" a]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   301
    unfolding o_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   302
  from pq have q: "Chi (of_rat \<circ> op ! q) = b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   303
    using euclidean_eq[of "Chi (of_rat \<circ> op ! q)" b]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   304
    unfolding o_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   305
  have "x \<in> ?box (p, q)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   306
    using p q ab by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   307
  thus "x \<in> UNION ?idx ?box" using ab e p q exI[of _ p] exI[of _ q] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   308
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   309
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   310
lemma borel_sigma_sets_subset:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   311
  "A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   312
  using sigma_sets_subset[of A borel] by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   313
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   314
lemma borel_eq_sigmaI1:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   315
  fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   316
  assumes borel_eq: "borel = sigma UNIV X"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   317
  assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (range F))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   318
  assumes F: "\<And>i. F i \<in> sets borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   319
  shows "borel = sigma UNIV (range F)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   320
  unfolding borel_def
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   321
proof (intro sigma_eqI antisym)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   322
  have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   323
    unfolding borel_def by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   324
  also have "\<dots> = sigma_sets UNIV X"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   325
    unfolding borel_eq by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   326
  also have "\<dots> \<subseteq> sigma_sets UNIV (range F)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   327
    using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   328
  finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (range F)" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   329
  show "sigma_sets UNIV (range F) \<subseteq> sigma_sets UNIV {S. open S}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   330
    unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   331
qed auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   332
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   333
lemma borel_eq_sigmaI2:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   334
  fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   335
    and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   336
  assumes borel_eq: "borel = sigma UNIV (range (\<lambda>(i, j). G i j))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   337
  assumes X: "\<And>i j. G i j \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   338
  assumes F: "\<And>i j. F i j \<in> sets borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   339
  shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   340
  using assms by (intro borel_eq_sigmaI1[where X="range (\<lambda>(i, j). G i j)" and F="(\<lambda>(i, j). F i j)"]) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   341
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   342
lemma borel_eq_sigmaI3:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   343
  fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   344
  assumes borel_eq: "borel = sigma UNIV X"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   345
  assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   346
  assumes F: "\<And>i j. F i j \<in> sets borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   347
  shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   348
  using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   349
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   350
lemma borel_eq_sigmaI4:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   351
  fixes F :: "'i \<Rightarrow> 'a::topological_space set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   352
    and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   353
  assumes borel_eq: "borel = sigma UNIV (range (\<lambda>(i, j). G i j))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   354
  assumes X: "\<And>i j. G i j \<in> sets (sigma UNIV (range F))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   355
  assumes F: "\<And>i. F i \<in> sets borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   356
  shows "borel = sigma UNIV (range F)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   357
  using assms by (intro borel_eq_sigmaI1[where X="range (\<lambda>(i, j). G i j)" and F=F]) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   358
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   359
lemma borel_eq_sigmaI5:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   360
  fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   361
  assumes borel_eq: "borel = sigma UNIV (range G)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   362
  assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   363
  assumes F: "\<And>i j. F i j \<in> sets borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   364
  shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   365
  using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   366
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   367
lemma halfspace_gt_in_halfspace:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   368
  "{x\<Colon>'a. a < x $$ i} \<in> sigma_sets UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   369
  (is "?set \<in> ?SIGMA")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   370
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   371
  interpret sigma_algebra UNIV ?SIGMA
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   372
    by (intro sigma_algebra_sigma_sets) simp_all
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   373
  have *: "?set = (\<Union>n. UNIV - {x\<Colon>'a. x $$ i < a + 1 / real (Suc n)})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   374
  proof (safe, simp_all add: not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   375
    fix x assume "a < x $$ i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   376
    with reals_Archimedean[of "x $$ i - a"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   377
    obtain n where "a + 1 / real (Suc n) < x $$ i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   378
      by (auto simp: inverse_eq_divide field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   379
    then show "\<exists>n. a + 1 / real (Suc n) \<le> x $$ i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   380
      by (blast intro: less_imp_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   381
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   382
    fix x n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   383
    have "a < a + 1 / real (Suc n)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   384
    also assume "\<dots> \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   385
    finally show "a < x" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   386
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   387
  show "?set \<in> ?SIGMA" unfolding *
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   388
    by (auto intro!: Diff)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   389
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   390
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   391
lemma borel_eq_halfspace_less:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   392
  "borel = sigma UNIV (range (\<lambda>(a, i). {x::'a::ordered_euclidean_space. x $$ i < a}))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   393
  (is "_ = ?SIGMA")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   394
proof (rule borel_eq_sigmaI3[OF borel_def])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   395
  fix S :: "'a set" assume "S \<in> {S. open S}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   396
  then have "open S" by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   397
  from open_UNION[OF this]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   398
  obtain I where *: "S =
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   399
    (\<Union>(a, b)\<in>I.
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   400
        (\<Inter> i<DIM('a). {x. (Chi (real_of_rat \<circ> op ! a)::'a) $$ i < x $$ i}) \<inter>
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   401
        (\<Inter> i<DIM('a). {x. x $$ i < (Chi (real_of_rat \<circ> op ! b)::'a) $$ i}))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   402
    unfolding greaterThanLessThan_def
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   403
    unfolding eucl_greaterThan_eq_halfspaces[where 'a='a]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   404
    unfolding eucl_lessThan_eq_halfspaces[where 'a='a]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   405
    by blast
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   406
  show "S \<in> ?SIGMA"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   407
    unfolding *
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   408
    by (safe intro!: countable_UN Int countable_INT) (auto intro!: halfspace_gt_in_halfspace)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   409
qed auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   410
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   411
lemma borel_eq_halfspace_le:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   412
  "borel = sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x $$ i \<le> a}))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   413
  (is "_ = ?SIGMA")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   414
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   415
  fix a i
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   416
  have *: "{x::'a. x$$i < a} = (\<Union>n. {x. x$$i \<le> a - 1/real (Suc n)})"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   417
  proof (safe, simp_all)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   418
    fix x::'a assume *: "x$$i < a"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   419
    with reals_Archimedean[of "a - x$$i"]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   420
    obtain n where "x $$ i < a - 1 / (real (Suc n))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   421
      by (auto simp: field_simps inverse_eq_divide)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   422
    then show "\<exists>n. x $$ i \<le> a - 1 / (real (Suc n))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   423
      by (blast intro: less_imp_le)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   424
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   425
    fix x::'a and n
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   426
    assume "x$$i \<le> a - 1 / real (Suc n)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   427
    also have "\<dots> < a" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   428
    finally show "x$$i < a" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   429
  qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   430
  show "{x. x$$i < a} \<in> ?SIGMA" unfolding *
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   431
    by (safe intro!: countable_UN) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   432
qed auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   433
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   434
lemma borel_eq_halfspace_ge:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   435
  "borel = sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a \<le> x $$ i}))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   436
  (is "_ = ?SIGMA")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   437
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   438
  fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \<le> x$$i}" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   439
  show "{x. x$$i < a} \<in> ?SIGMA" unfolding *
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   440
      by (safe intro!: compl_sets) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   441
qed auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   442
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   443
lemma borel_eq_halfspace_greater:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   444
  "borel = sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a < x $$ i}))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   445
  (is "_ = ?SIGMA")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   446
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   447
  fix a i have *: "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   448
  show "{x. x$$i \<le> a} \<in> ?SIGMA" unfolding *
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   449
    by (safe intro!: compl_sets) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   450
qed auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   451
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   452
lemma borel_eq_atMost:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   453
  "borel = sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   454
  (is "_ = ?SIGMA")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   455
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   456
  fix a i show "{x. x$$i \<le> a} \<in> ?SIGMA"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   457
  proof cases
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   458
    assume "i < DIM('a)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   459
    then have *: "{x::'a. x$$i \<le> a} = (\<Union>k::nat. {.. (\<chi>\<chi> n. if n = i then a else real k)})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   460
    proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   461
      fix x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   462
      from real_arch_simple[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"] guess k::nat ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   463
      then have "\<And>i. i < DIM('a) \<Longrightarrow> x$$i \<le> real k"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   464
        by (subst (asm) Max_le_iff) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   465
      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia \<le> real k"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   466
        by (auto intro!: exI[of _ k])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   467
    qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   468
    show "{x. x$$i \<le> a} \<in> ?SIGMA" unfolding *
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   469
      by (safe intro!: countable_UN) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   470
  qed (auto intro: sigma_sets_top sigma_sets.Empty)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   471
qed auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   472
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   473
lemma borel_eq_greaterThan:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   474
  "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {a<..}))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   475
  (is "_ = ?SIGMA")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   476
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   477
  fix a i show "{x. x$$i \<le> a} \<in> ?SIGMA"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   478
  proof cases
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   479
    assume "i < DIM('a)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   480
    have "{x::'a. x$$i \<le> a} = UNIV - {x::'a. a < x$$i}" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   481
    also have *: "{x::'a. a < x$$i} = (\<Union>k::nat. {(\<chi>\<chi> n. if n = i then a else -real k) <..})" using `i <DIM('a)`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   482
    proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   483
      fix x
44666
8670a39d4420 remove more duplicate lemmas
huffman
parents: 44537
diff changeset
   484
      from reals_Archimedean2[of "Max ((\<lambda>i. -x$$i)`{..<DIM('a)})"]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   485
      guess k::nat .. note k = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   486
      { fix i assume "i < DIM('a)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   487
        then have "-x$$i < real k"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   488
          using k by (subst (asm) Max_less_iff) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   489
        then have "- real k < x$$i" by simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   490
      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> -real k < x $$ ia"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   491
        by (auto intro!: exI[of _ k])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   492
    qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   493
    finally show "{x. x$$i \<le> a} \<in> ?SIGMA"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   494
      apply (simp only:)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   495
      apply (safe intro!: countable_UN Diff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   496
      apply (auto intro: sigma_sets_top)
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45288
diff changeset
   497
      done
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   498
  qed (auto intro: sigma_sets_top sigma_sets.Empty)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   499
qed auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   500
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   501
lemma borel_eq_lessThan:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   502
  "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {..<a}))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   503
  (is "_ = ?SIGMA")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   504
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   505
  fix a i show "{x. a \<le> x$$i} \<in> ?SIGMA"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   506
  proof cases
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   507
    fix a i assume "i < DIM('a)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   508
    have "{x::'a. a \<le> x$$i} = UNIV - {x::'a. x$$i < a}" by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   509
    also have *: "{x::'a. x$$i < a} = (\<Union>k::nat. {..< (\<chi>\<chi> n. if n = i then a else real k)})" using `i <DIM('a)`
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   510
    proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   511
      fix x
44666
8670a39d4420 remove more duplicate lemmas
huffman
parents: 44537
diff changeset
   512
      from reals_Archimedean2[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   513
      guess k::nat .. note k = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   514
      { fix i assume "i < DIM('a)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   515
        then have "x$$i < real k"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   516
          using k by (subst (asm) Max_less_iff) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   517
        then have "x$$i < real k" by simp }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   518
      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia < real k"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   519
        by (auto intro!: exI[of _ k])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   520
    qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   521
    finally show "{x. a \<le> x$$i} \<in> ?SIGMA"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   522
      apply (simp only:)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   523
      apply (safe intro!: countable_UN Diff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   524
      apply (auto intro: sigma_sets_top)
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45288
diff changeset
   525
      done
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   526
  qed (auto intro: sigma_sets_top sigma_sets.Empty)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   527
qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   528
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   529
lemma borel_eq_atLeastAtMost:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   530
  "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} \<Colon>'a\<Colon>ordered_euclidean_space set))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   531
  (is "_ = ?SIGMA")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   532
proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   533
  fix a::'a
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   534
  have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   535
  proof (safe, simp_all add: eucl_le[where 'a='a])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   536
    fix x
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   537
    from real_arch_simple[of "Max ((\<lambda>i. - x$$i)`{..<DIM('a)})"]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   538
    guess k::nat .. note k = this
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   539
    { fix i assume "i < DIM('a)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   540
      with k have "- x$$i \<le> real k"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   541
        by (subst (asm) Max_le_iff) (auto simp: field_simps)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   542
      then have "- real k \<le> x$$i" by simp }
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   543
    then show "\<exists>n::nat. \<forall>i<DIM('a). - real n \<le> x $$ i"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   544
      by (auto intro!: exI[of _ k])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   545
  qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   546
  show "{..a} \<in> ?SIGMA" unfolding *
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   547
    by (safe intro!: countable_UN)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   548
       (auto intro!: sigma_sets_top)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   549
qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   550
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   551
lemma borel_eq_greaterThanLessThan:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   552
  "borel = sigma UNIV (range (\<lambda> (a, b). {a <..< b} :: 'a \<Colon> ordered_euclidean_space set))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   553
    (is "_ = ?SIGMA")
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   554
proof (rule borel_eq_sigmaI1[OF borel_def])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   555
  fix M :: "'a set" assume "M \<in> {S. open S}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   556
  then have "open M" by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   557
  show "M \<in> ?SIGMA"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   558
    apply (subst open_UNION[OF `open M`])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   559
    apply (safe intro!: countable_UN)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   560
    apply auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   561
    done
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   562
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   563
42862
7d7627738e66 add borel_eq_atLeastLessThan
hoelzl
parents: 42150
diff changeset
   564
lemma borel_eq_atLeastLessThan:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   565
  "borel = sigma UNIV (range (\<lambda>(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   566
proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   567
  have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   568
  fix x :: real
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   569
  have "{..<x} = (\<Union>i::nat. {-real i ..< x})"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   570
    by (auto simp: move_uminus real_arch_simple)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   571
  then show "{..< x} \<in> ?SIGMA"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   572
    by (auto intro: sigma_sets.intros)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   573
qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   574
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   575
lemma borel_measurable_halfspacesI:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   576
  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   577
  assumes F: "borel = sigma UNIV (range F)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   578
  and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M" 
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   579
  and S: "\<And>a i. \<not> i < DIM('c) \<Longrightarrow> S a i \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   580
  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a::real. S a i \<in> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   581
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   582
  fix a :: real and i assume i: "i < DIM('c)" and f: "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   583
  then show "S a i \<in> sets M" unfolding assms
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   584
    by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1))
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   585
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   586
  assume a: "\<forall>i<DIM('c). \<forall>a. S a i \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   587
  { fix a i have "S a i \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   588
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   589
      assume "i < DIM('c)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   590
      with a show ?thesis unfolding assms(2) by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   591
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   592
      assume "\<not> i < DIM('c)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   593
      from S[OF this] show ?thesis .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   594
    qed }
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   595
  then show "f \<in> borel_measurable M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   596
    by (auto intro!: measurable_measure_of simp: S_eq F)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   597
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   598
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   599
lemma borel_measurable_iff_halfspace_le:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   600
  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   601
  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i \<le> a} \<in> sets M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   602
  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   603
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   604
lemma borel_measurable_iff_halfspace_less:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   605
  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   606
  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i < a} \<in> sets M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   607
  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   608
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   609
lemma borel_measurable_iff_halfspace_ge:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   610
  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   611
  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a \<le> f w $$ i} \<in> sets M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   612
  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   613
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   614
lemma borel_measurable_iff_halfspace_greater:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   615
  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   616
  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a < f w $$ i} \<in> sets M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   617
  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   618
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   619
lemma borel_measurable_iff_le:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   620
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   621
  using borel_measurable_iff_halfspace_le[where 'c=real] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   622
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   623
lemma borel_measurable_iff_less:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   624
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   625
  using borel_measurable_iff_halfspace_less[where 'c=real] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   626
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   627
lemma borel_measurable_iff_ge:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   628
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   629
  using borel_measurable_iff_halfspace_ge[where 'c=real] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   630
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   631
lemma borel_measurable_iff_greater:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   632
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   633
  using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   634
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   635
lemma borel_measurable_euclidean_component':
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   636
  "(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   637
proof (rule borel_measurableI)
44537
c10485a6a7af make HOL-Probability respect set/pred distinction
huffman
parents: 44282
diff changeset
   638
  fix S::"real set" assume "open S"
39087
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   639
  from open_vimage_euclidean_component[OF this]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   640
  show "(\<lambda>x. x $$ i) -` S \<inter> space borel \<in> sets borel"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   641
    by (auto intro: borel_open)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   642
qed
39087
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   643
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   644
lemma borel_measurable_euclidean_component:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   645
  fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   646
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   647
  shows "(\<lambda>x. f x $$ i) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   648
  using measurable_comp[OF f borel_measurable_euclidean_component'] by (simp add: comp_def)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   649
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   650
lemma borel_measurable_euclidean_space:
39087
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   651
  fixes f :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   652
  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M)"
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   653
proof safe
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   654
  fix i assume "f \<in> borel_measurable M"
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   655
  then show "(\<lambda>x. f x $$ i) \<in> borel_measurable M"
41025
8b2cd85ecf11 fixed spelling errors
hoelzl
parents: 41023
diff changeset
   656
    by (auto intro: borel_measurable_euclidean_component)
39087
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   657
next
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   658
  assume f: "\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M"
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   659
  then show "f \<in> borel_measurable M"
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   660
    unfolding borel_measurable_iff_halfspace_le by auto
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   661
qed
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   662
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   663
subsection "Borel measurable operators"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   664
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   665
lemma borel_measurable_continuous_on1:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   666
  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   667
  assumes "continuous_on UNIV f"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   668
  shows "f \<in> borel_measurable borel"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   669
  apply(rule borel_measurableI)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   670
  using continuous_open_preimage[OF assms] unfolding vimage_def by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   671
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   672
lemma borel_measurable_continuous_on:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   673
  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   674
  assumes f: "continuous_on UNIV f" and g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   675
  shows "(\<lambda>x. f (g x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   676
  using measurable_comp[OF g borel_measurable_continuous_on1[OF f]] by (simp add: comp_def)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   677
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   678
lemma borel_measurable_continuous_on_open':
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   679
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   680
  assumes cont: "continuous_on A f" "open A"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   681
  shows "(\<lambda>x. if x \<in> A then f x else c) \<in> borel_measurable borel" (is "?f \<in> _")
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   682
proof (rule borel_measurableI)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   683
  fix S :: "'b set" assume "open S"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   684
  then have "open {x\<in>A. f x \<in> S}"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   685
    by (intro continuous_open_preimage[OF cont]) auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   686
  then have *: "{x\<in>A. f x \<in> S} \<in> sets borel" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   687
  have "?f -` S \<inter> space borel = 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   688
    {x\<in>A. f x \<in> S} \<union> (if c \<in> S then space borel - A else {})"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   689
    by (auto split: split_if_asm)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   690
  also have "\<dots> \<in> sets borel"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   691
    using * `open A` by (auto simp del: space_borel intro!: Un)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   692
  finally show "?f -` S \<inter> space borel \<in> sets borel" .
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   693
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   694
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   695
lemma borel_measurable_continuous_on_open:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   696
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   697
  assumes cont: "continuous_on A f" "open A"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   698
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   699
  shows "(\<lambda>x. if g x \<in> A then f (g x) else c) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   700
  using measurable_comp[OF g borel_measurable_continuous_on_open'[OF cont], of c]
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   701
  by (simp add: comp_def)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   702
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   703
lemma borel_measurable_uminus[simp, intro]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   704
  fixes g :: "'a \<Rightarrow> real"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   705
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   706
  shows "(\<lambda>x. - g x) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   707
  by (rule borel_measurable_continuous_on[OF _ g]) (auto intro: continuous_on_minus continuous_on_id)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   708
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   709
lemma euclidean_component_prod:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   710
  fixes x :: "'a :: euclidean_space \<times> 'b :: euclidean_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   711
  shows "x $$ i = (if i < DIM('a) then fst x $$ i else snd x $$ (i - DIM('a)))"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   712
  unfolding euclidean_component_def basis_prod_def inner_prod_def by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   713
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   714
lemma borel_measurable_Pair[simp, intro]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   715
  fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space" and g :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   716
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   717
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   718
  shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   719
proof (intro borel_measurable_iff_halfspace_le[THEN iffD2] allI impI)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   720
  fix i and a :: real assume i: "i < DIM('b \<times> 'c)"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   721
  have [simp]: "\<And>P A B C. {w. (P \<longrightarrow> A w \<and> B w) \<and> (\<not> P \<longrightarrow> A w \<and> C w)} = 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   722
    {w. A w \<and> (P \<longrightarrow> B w) \<and> (\<not> P \<longrightarrow> C w)}" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   723
  from i f g show "{w \<in> space M. (f w, g w) $$ i \<le> a} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   724
    by (auto simp: euclidean_component_prod intro!: sets_Collect borel_measurable_euclidean_component)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   725
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   726
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   727
lemma continuous_on_fst: "continuous_on UNIV fst"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   728
proof -
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   729
  have [simp]: "range fst = UNIV" by (auto simp: image_iff)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   730
  show ?thesis
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   731
    using closed_vimage_fst
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   732
    by (auto simp: continuous_on_closed closed_closedin vimage_def)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   733
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   734
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   735
lemma continuous_on_snd: "continuous_on UNIV snd"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   736
proof -
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   737
  have [simp]: "range snd = UNIV" by (auto simp: image_iff)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   738
  show ?thesis
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   739
    using closed_vimage_snd
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   740
    by (auto simp: continuous_on_closed closed_closedin vimage_def)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   741
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   742
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   743
lemma borel_measurable_continuous_Pair:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   744
  fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space" and g :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   745
  assumes [simp]: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   746
  assumes [simp]: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   747
  assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   748
  shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   749
proof -
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   750
  have eq: "(\<lambda>x. H (f x) (g x)) = (\<lambda>x. (\<lambda>x. H (fst x) (snd x)) (f x, g x))" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   751
  show ?thesis
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   752
    unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   753
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   754
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   755
lemma borel_measurable_add[simp, intro]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   756
  fixes f g :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   757
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   758
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   759
  shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   760
  using f g
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   761
  by (rule borel_measurable_continuous_Pair)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   762
     (auto intro: continuous_on_fst continuous_on_snd continuous_on_add)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   763
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   764
lemma borel_measurable_setsum[simp, intro]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   765
  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   766
  assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   767
  shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   768
proof cases
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   769
  assume "finite S"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   770
  thus ?thesis using assms by induct auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   771
qed simp
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   772
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   773
lemma borel_measurable_diff[simp, intro]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   774
  fixes f :: "'a \<Rightarrow> real"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   775
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   776
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   777
  shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   778
  unfolding diff_minus using assms by fast
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   779
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   780
lemma borel_measurable_times[simp, intro]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   781
  fixes f :: "'a \<Rightarrow> real"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   782
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   783
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   784
  shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   785
  using f g
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   786
  by (rule borel_measurable_continuous_Pair)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   787
     (auto intro: continuous_on_fst continuous_on_snd continuous_on_mult)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   788
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   789
lemma continuous_on_dist:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   790
  fixes f :: "'a :: t2_space \<Rightarrow> 'b :: metric_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   791
  shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. dist (f x) (g x))"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   792
  unfolding continuous_on_eq_continuous_within by (auto simp: continuous_dist)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   793
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   794
lemma borel_measurable_dist[simp, intro]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   795
  fixes g f :: "'a \<Rightarrow> 'b::ordered_euclidean_space"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   796
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   797
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   798
  shows "(\<lambda>x. dist (f x) (g x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   799
  using f g
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   800
  by (rule borel_measurable_continuous_Pair)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   801
     (intro continuous_on_dist continuous_on_fst continuous_on_snd)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   802
  
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   803
lemma affine_borel_measurable_vector:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   804
  fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   805
  assumes "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   806
  shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   807
proof (rule borel_measurableI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   808
  fix S :: "'x set" assume "open S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   809
  show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   810
  proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   811
    assume "b \<noteq> 0"
44537
c10485a6a7af make HOL-Probability respect set/pred distinction
huffman
parents: 44282
diff changeset
   812
    with `open S` have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S")
c10485a6a7af make HOL-Probability respect set/pred distinction
huffman
parents: 44282
diff changeset
   813
      by (auto intro!: open_affinity simp: scaleR_add_right)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   814
    hence "?S \<in> sets borel" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   815
    moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   816
    from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   817
      apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   818
    ultimately show ?thesis using assms unfolding in_borel_measurable_borel
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   819
      by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   820
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   821
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   822
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   823
lemma affine_borel_measurable:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   824
  fixes g :: "'a \<Rightarrow> real"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   825
  assumes g: "g \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   826
  shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   827
  using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   828
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   829
lemma borel_measurable_setprod[simp, intro]:
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   830
  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   831
  assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   832
  shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   833
proof cases
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   834
  assume "finite S"
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   835
  thus ?thesis using assms by induct auto
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   836
qed simp
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   837
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   838
lemma borel_measurable_inverse[simp, intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   839
  fixes f :: "'a \<Rightarrow> real"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   840
  assumes f: "f \<in> borel_measurable M"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   841
  shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   842
proof -
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   843
  have *: "\<And>x::real. inverse x = (if x \<in> UNIV - {0} then inverse x else 0)" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   844
  show ?thesis
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   845
    apply (subst *)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   846
    apply (rule borel_measurable_continuous_on_open)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   847
    apply (auto intro!: f continuous_on_inverse continuous_on_id)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   848
    done
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   849
qed
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   850
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   851
lemma borel_measurable_divide[simp, intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   852
  fixes f :: "'a \<Rightarrow> real"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   853
  assumes "f \<in> borel_measurable M"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   854
  and "g \<in> borel_measurable M"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   855
  shows "(\<lambda>x. f x / g x) \<in> borel_measurable M"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   856
  unfolding field_divide_inverse
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   857
  by (rule borel_measurable_inverse borel_measurable_times assms)+
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   858
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   859
lemma borel_measurable_max[intro, simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   860
  fixes f g :: "'a \<Rightarrow> real"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   861
  assumes "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   862
  assumes "g \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   863
  shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   864
  unfolding max_def by (auto intro!: assms measurable_If)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   865
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   866
lemma borel_measurable_min[intro, simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   867
  fixes f g :: "'a \<Rightarrow> real"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   868
  assumes "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   869
  assumes "g \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   870
  shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   871
  unfolding min_def by (auto intro!: assms measurable_If)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   872
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   873
lemma borel_measurable_abs[simp, intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   874
  assumes "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   875
  shows "(\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   876
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   877
  have *: "\<And>x. \<bar>f x\<bar> = max 0 (f x) + max 0 (- f x)" by (simp add: max_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   878
  show ?thesis unfolding * using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   879
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   880
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   881
lemma borel_measurable_nth[simp, intro]:
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   882
  "(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   883
  using borel_measurable_euclidean_component'
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   884
  unfolding nth_conv_component by auto
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   885
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   886
lemma convex_measurable:
42990
3706951a6421 composition of convex and measurable function is measurable
hoelzl
parents: 42950
diff changeset
   887
  fixes a b :: real
3706951a6421 composition of convex and measurable function is measurable
hoelzl
parents: 42950
diff changeset
   888
  assumes X: "X \<in> borel_measurable M" "X ` space M \<subseteq> { a <..< b}"
3706951a6421 composition of convex and measurable function is measurable
hoelzl
parents: 42950
diff changeset
   889
  assumes q: "convex_on { a <..< b} q"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   890
  shows "(\<lambda>x. q (X x)) \<in> borel_measurable M"
42990
3706951a6421 composition of convex and measurable function is measurable
hoelzl
parents: 42950
diff changeset
   891
proof -
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   892
  have "(\<lambda>x. if X x \<in> {a <..< b} then q (X x) else 0) \<in> borel_measurable M" (is "?qX")
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   893
  proof (rule borel_measurable_continuous_on_open[OF _ _ X(1)])
42990
3706951a6421 composition of convex and measurable function is measurable
hoelzl
parents: 42950
diff changeset
   894
    show "open {a<..<b}" by auto
3706951a6421 composition of convex and measurable function is measurable
hoelzl
parents: 42950
diff changeset
   895
    from this q show "continuous_on {a<..<b} q"
3706951a6421 composition of convex and measurable function is measurable
hoelzl
parents: 42950
diff changeset
   896
      by (rule convex_on_continuous)
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   897
  qed
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   898
  moreover have "?qX \<longleftrightarrow> (\<lambda>x. q (X x)) \<in> borel_measurable M"
42990
3706951a6421 composition of convex and measurable function is measurable
hoelzl
parents: 42950
diff changeset
   899
    using X by (intro measurable_cong) auto
3706951a6421 composition of convex and measurable function is measurable
hoelzl
parents: 42950
diff changeset
   900
  ultimately show ?thesis by simp
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   901
qed
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   902
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   903
lemma borel_measurable_ln[simp,intro]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   904
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   905
  shows "(\<lambda>x. ln (f x)) \<in> borel_measurable M"
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   906
proof -
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   907
  { fix x :: real assume x: "x \<le> 0"
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   908
    { fix x::real assume "x \<le> 0" then have "\<And>u. exp u = x \<longleftrightarrow> False" by auto }
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   909
    from this[of x] x this[of 0] have "ln 0 = ln x"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   910
      by (auto simp: ln_def) }
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   911
  note ln_imp = this
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   912
  have "(\<lambda>x. if f x \<in> {0<..} then ln (f x) else ln 0) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   913
  proof (rule borel_measurable_continuous_on_open[OF _ _ f])
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   914
    show "continuous_on {0<..} ln"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   915
      by (auto intro!: continuous_at_imp_continuous_on DERIV_ln DERIV_isCont
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   916
               simp: continuous_isCont[symmetric])
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   917
    show "open ({0<..}::real set)" by auto
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   918
  qed
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   919
  also have "(\<lambda>x. if x \<in> {0<..} then ln x else ln 0) = ln"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   920
    by (simp add: fun_eq_iff not_less ln_imp)
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   921
  finally show ?thesis .
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   922
qed
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   923
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   924
lemma borel_measurable_log[simp,intro]:
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   925
  "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log b (f x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   926
  unfolding log_def by auto
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   927
47761
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   928
lemma borel_measurable_real_floor:
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   929
  "(\<lambda>x::real. real \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   930
  unfolding borel_measurable_iff_ge
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   931
proof (intro allI)
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   932
  fix a :: real
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   933
  { fix x have "a \<le> real \<lfloor>x\<rfloor> \<longleftrightarrow> real \<lceil>a\<rceil> \<le> x"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   934
      using le_floor_eq[of "\<lceil>a\<rceil>" x] ceiling_le_iff[of a "\<lfloor>x\<rfloor>"]
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   935
      unfolding real_eq_of_int by simp }
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   936
  then have "{w::real \<in> space borel. a \<le> real \<lfloor>w\<rfloor>} = {real \<lceil>a\<rceil>..}" by auto
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   937
  then show "{w::real \<in> space borel. a \<le> real \<lfloor>w\<rfloor>} \<in> sets borel" by auto
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   938
qed
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   939
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   940
lemma borel_measurable_real_natfloor[intro, simp]:
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   941
  assumes "f \<in> borel_measurable M"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   942
  shows "(\<lambda>x. real (natfloor (f x))) \<in> borel_measurable M"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   943
proof -
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   944
  have "\<And>x. real (natfloor (f x)) = max 0 (real \<lfloor>f x\<rfloor>)"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   945
    by (auto simp: max_def natfloor_def)
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   946
  with borel_measurable_max[OF measurable_comp[OF assms borel_measurable_real_floor] borel_measurable_const]
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   947
  show ?thesis by (simp add: comp_def)
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   948
qed
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   949
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
   950
subsection "Borel space on the extended reals"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
   951
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   952
lemma borel_measurable_ereal[simp, intro]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
   953
  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   954
  using continuous_on_ereal f by (rule borel_measurable_continuous_on)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
   955
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   956
lemma borel_measurable_real_of_ereal[simp, intro]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   957
  fixes f :: "'a \<Rightarrow> ereal" 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   958
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   959
  shows "(\<lambda>x. real (f x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   960
proof -
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   961
  have "(\<lambda>x. if f x \<in> UNIV - { \<infinity>, - \<infinity> } then real (f x) else 0) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   962
    using continuous_on_real
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   963
    by (rule borel_measurable_continuous_on_open[OF _ _ f]) auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   964
  also have "(\<lambda>x. if f x \<in> UNIV - { \<infinity>, - \<infinity> } then real (f x) else 0) = (\<lambda>x. real (f x))"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   965
    by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   966
  finally show ?thesis .
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   967
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   968
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   969
lemma borel_measurable_ereal_cases:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   970
  fixes f :: "'a \<Rightarrow> ereal" 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   971
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   972
  assumes H: "(\<lambda>x. H (ereal (real (f x)))) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   973
  shows "(\<lambda>x. H (f x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   974
proof -
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   975
  let ?F = "\<lambda>x. if x \<in> f -` {\<infinity>} then H \<infinity> else if x \<in> f -` {-\<infinity>} then H (-\<infinity>) else H (ereal (real (f x)))"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   976
  { fix x have "H (f x) = ?F x" by (cases "f x") auto }
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   977
  moreover 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   978
  have "?F \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   979
    by (intro measurable_If_set f measurable_sets[OF f] H) auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   980
  ultimately
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   981
  show ?thesis by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   982
qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
   983
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   984
lemma
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   985
  fixes f :: "'a \<Rightarrow> ereal" assumes f[simp]: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   986
  shows borel_measurable_ereal_abs[intro, simp]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   987
    and borel_measurable_ereal_inverse[simp, intro]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   988
    and borel_measurable_uminus_ereal[intro]: "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   989
  by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   990
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   991
lemma borel_measurable_uminus_eq_ereal[simp]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   992
  "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   993
proof
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   994
  assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   995
qed auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   996
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   997
lemma sets_Collect_If_set:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   998
  assumes "A \<inter> space M \<in> sets M" "{x \<in> space M. P x} \<in> sets M" "{x \<in> space M. Q x} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   999
  shows "{x \<in> space M. if x \<in> A then P x else Q x} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1000
proof -
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1001
  have *: "{x \<in> space M. if x \<in> A then P x else Q x} = 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1002
    {x \<in> space M. if x \<in> A \<inter> space M then P x else Q x}" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1003
  show ?thesis unfolding * unfolding if_bool_eq_conj using assms
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1004
    by (auto intro!: sets_Collect simp: Int_def conj_commute)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1005
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1006
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1007
lemma set_Collect_ereal2:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1008
  fixes f g :: "'a \<Rightarrow> ereal" 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1009
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1010
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1011
  assumes H: "{x \<in> space M. H (ereal (real (f x))) (ereal (real (g x)))} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1012
    "{x \<in> space M. H (-\<infinity>) (ereal (real (g x)))} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1013
    "{x \<in> space M. H (\<infinity>) (ereal (real (g x)))} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1014
    "{x \<in> space M. H (ereal (real (f x))) (-\<infinity>)} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1015
    "{x \<in> space M. H (ereal (real (f x))) (\<infinity>)} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1016
  shows "{x \<in> space M. H (f x) (g x)} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1017
proof -
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1018
  let ?G = "\<lambda>y x. if x \<in> g -` {\<infinity>} then H y \<infinity> else if x \<in> g -` {-\<infinity>} then H y (-\<infinity>) else H y (ereal (real (g x)))"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1019
  let ?F = "\<lambda>x. if x \<in> f -` {\<infinity>} then ?G \<infinity> x else if x \<in> f -` {-\<infinity>} then ?G (-\<infinity>) x else ?G (ereal (real (f x))) x"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1020
  { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1021
  moreover 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1022
  have "{x \<in> space M. ?F x} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1023
    by (intro sets_Collect H measurable_sets[OF f] measurable_sets[OF g] sets_Collect_If_set) auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1024
  ultimately
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1025
  show ?thesis by simp
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1026
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1027
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1028
lemma
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1029
  fixes f g :: "'a \<Rightarrow> ereal"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1030
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1031
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1032
  shows borel_measurable_ereal_le[intro,simp]: "{x \<in> space M. f x \<le> g x} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1033
    and borel_measurable_ereal_less[intro,simp]: "{x \<in> space M. f x < g x} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1034
    and borel_measurable_ereal_eq[intro,simp]: "{w \<in> space M. f w = g w} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1035
    and borel_measurable_ereal_neq[intro,simp]: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1036
  using f g by (auto simp: f g set_Collect_ereal2[OF f g] intro!: sets_Collect_neg)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1037
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1038
lemma borel_measurable_ereal_iff:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1039
  shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1040
proof
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1041
  assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1042
  from borel_measurable_real_of_ereal[OF this]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1043
  show "f \<in> borel_measurable M" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1044
qed auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1045
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1046
lemma borel_measurable_ereal_iff_real:
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1047
  fixes f :: "'a \<Rightarrow> ereal"
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1048
  shows "f \<in> borel_measurable M \<longleftrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1049
    ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1050
proof safe
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1051
  assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<inter> space M \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1052
  have "f -` {\<infinity>} \<inter> space M = {x\<in>space M. f x = \<infinity>}" "f -` {-\<infinity>} \<inter> space M = {x\<in>space M. f x = -\<infinity>}" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1053
  with * have **: "{x\<in>space M. f x = \<infinity>} \<in> sets M" "{x\<in>space M. f x = -\<infinity>} \<in> sets M" by simp_all
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45288
diff changeset
  1054
  let ?f = "\<lambda>x. if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else ereal (real (f x))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1055
  have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1056
  also have "?f = f" by (auto simp: fun_eq_iff ereal_real)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1057
  finally show "f \<in> borel_measurable M" .
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1058
qed (auto intro: measurable_sets borel_measurable_real_of_ereal)
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
  1059
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1060
lemma borel_measurable_eq_atMost_ereal:
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1061
  fixes f :: "'a \<Rightarrow> ereal"
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1062
  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1063
proof (intro iffI allI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1064
  assume pos[rule_format]: "\<forall>a. f -` {..a} \<inter> space M \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1065
  show "f \<in> borel_measurable M"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1066
    unfolding borel_measurable_ereal_iff_real borel_measurable_iff_le
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1067
  proof (intro conjI allI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1068
    fix a :: real
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1069
    { fix x :: ereal assume *: "\<forall>i::nat. real i < x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1070
      have "x = \<infinity>"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1071
      proof (rule ereal_top)
44666
8670a39d4420 remove more duplicate lemmas
huffman
parents: 44537
diff changeset
  1072
        fix B from reals_Archimedean2[of B] guess n ..
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1073
        then have "ereal B < real n" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1074
        with * show "B \<le> x" by (metis less_trans less_imp_le)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1075
      qed }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1076
    then have "f -` {\<infinity>} \<inter> space M = space M - (\<Union>i::nat. f -` {.. real i} \<inter> space M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1077
      by (auto simp: not_le)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1078
    then show "f -` {\<infinity>} \<inter> space M \<in> sets M" using pos by (auto simp del: UN_simps intro!: Diff)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1079
    moreover
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1080
    have "{-\<infinity>::ereal} = {..-\<infinity>}" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1081
    then show "f -` {-\<infinity>} \<inter> space M \<in> sets M" using pos by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1082
    moreover have "{x\<in>space M. f x \<le> ereal a} \<in> sets M"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1083
      using pos[of "ereal a"] by (simp add: vimage_def Int_def conj_commute)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1084
    moreover have "{w \<in> space M. real (f w) \<le> a} =
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1085
      (if a < 0 then {w \<in> space M. f w \<le> ereal a} - f -` {-\<infinity>} \<inter> space M
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1086
      else {w \<in> space M. f w \<le> ereal a} \<union> (f -` {\<infinity>} \<inter> space M) \<union> (f -` {-\<infinity>} \<inter> space M))" (is "?l = ?r")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1087
      proof (intro set_eqI) fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases "f x") auto qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1088
    ultimately show "{w \<in> space M. real (f w) \<le> a} \<in> sets M" by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 35347
diff changeset
  1089
  qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1090
qed (simp add: measurable_sets)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 35347
diff changeset
  1091
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1092
lemma borel_measurable_eq_atLeast_ereal:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1093
  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1094
proof
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1095
  assume pos: "\<forall>a. f -` {a..} \<inter> space M \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1096
  moreover have "\<And>a. (\<lambda>x. - f x) -` {..a} = f -` {-a ..}"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1097
    by (auto simp: ereal_uminus_le_reorder)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1098
  ultimately have "(\<lambda>x. - f x) \<in> borel_measurable M"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1099
    unfolding borel_measurable_eq_atMost_ereal by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1100
  then show "f \<in> borel_measurable M" by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1101
qed (simp add: measurable_sets)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 35347
diff changeset
  1102
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1103
lemma greater_eq_le_measurable:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1104
  fixes f :: "'a \<Rightarrow> 'c::linorder"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1105
  shows "f -` {..< a} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {a ..} \<inter> space M \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1106
proof
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1107
  assume "f -` {a ..} \<inter> space M \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1108
  moreover have "f -` {..< a} \<inter> space M = space M - f -` {a ..} \<inter> space M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1109
  ultimately show "f -` {..< a} \<inter> space M \<in> sets M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1110
next
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1111
  assume "f -` {..< a} \<inter> space M \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1112
  moreover have "f -` {a ..} \<inter> space M = space M - f -` {..< a} \<inter> space M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1113
  ultimately show "f -` {a ..} \<inter> space M \<in> sets M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1114
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1115
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1116
lemma borel_measurable_ereal_iff_less:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1117
  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1118
  unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable ..
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1119
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1120
lemma less_eq_ge_measurable:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1121
  fixes f :: "'a \<Rightarrow> 'c::linorder"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1122
  shows "f -` {a <..} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {..a} \<inter> space M \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1123
proof
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1124
  assume "f -` {a <..} \<inter> space M \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1125
  moreover have "f -` {..a} \<inter> space M = space M - f -` {a <..} \<inter> space M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1126
  ultimately show "f -` {..a} \<inter> space M \<in> sets M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1127
next
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1128
  assume "f -` {..a} \<inter> space M \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1129
  moreover have "f -` {a <..} \<inter> space M = space M - f -` {..a} \<inter> space M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1130
  ultimately show "f -` {a <..} \<inter> space M \<in> sets M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1131
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1132
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1133
lemma borel_measurable_ereal_iff_ge:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1134
  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1135
  unfolding borel_measurable_eq_atMost_ereal less_eq_ge_measurable ..
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1136
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1137
lemma borel_measurable_ereal2:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1138
  fixes f g :: "'a \<Rightarrow> ereal" 
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1139
  assumes f: "f \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1140
  assumes g: "g \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1141
  assumes H: "(\<lambda>x. H (ereal (real (f x))) (ereal (real (g x)))) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1142
    "(\<lambda>x. H (-\<infinity>) (ereal (real (g x)))) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1143
    "(\<lambda>x. H (\<infinity>) (ereal (real (g x)))) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1144
    "(\<lambda>x. H (ereal (real (f x))) (-\<infinity>)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1145
    "(\<lambda>x. H (ereal (real (f x))) (\<infinity>)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1146
  shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1147
proof -
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1148
  let ?G = "\<lambda>y x. if x \<in> g -` {\<infinity>} then H y \<infinity> else if x \<in> g -` {-\<infinity>} then H y (-\<infinity>) else H y (ereal (real (g x)))"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1149
  let ?F = "\<lambda>x. if x \<in> f -` {\<infinity>} then ?G \<infinity> x else if x \<in> f -` {-\<infinity>} then ?G (-\<infinity>) x else ?G (ereal (real (f x))) x"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1150
  { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1151
  moreover 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1152
  have "?F \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1153
    by (intro measurable_If_set measurable_sets[OF f] measurable_sets[OF g] H) auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1154
  ultimately
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1155
  show ?thesis by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1156
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1157
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1158
lemma
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1159
  fixes f :: "'a \<Rightarrow> ereal" assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1160
  shows borel_measurable_ereal_eq_const: "{x\<in>space M. f x = c} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1161
    and borel_measurable_ereal_neq_const: "{x\<in>space M. f x \<noteq> c} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1162
  using f by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1163
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1164
lemma split_sets:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1165
  "{x\<in>space M. P x \<or> Q x} = {x\<in>space M. P x} \<union> {x\<in>space M. Q x}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1166
  "{x\<in>space M. P x \<and> Q x} = {x\<in>space M. P x} \<inter> {x\<in>space M. Q x}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1167
  by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1168
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1169
lemma
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1170
  fixes f :: "'a \<Rightarrow> ereal"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1171
  assumes [simp]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1172
  shows borel_measurable_ereal_add[intro, simp]: "(\<lambda>x. f x + g x) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1173
    and borel_measurable_ereal_times[intro, simp]: "(\<lambda>x. f x * g x) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1174
    and borel_measurable_ereal_min[simp, intro]: "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1175
    and borel_measurable_ereal_max[simp, intro]: "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1176
  by (auto simp add: borel_measurable_ereal2 measurable_If min_def max_def)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1177
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1178
lemma
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1179
  fixes f g :: "'a \<Rightarrow> ereal"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1180
  assumes "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1181
  assumes "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1182
  shows borel_measurable_ereal_diff[simp, intro]: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1183
    and borel_measurable_ereal_divide[simp, intro]: "(\<lambda>x. f x / g x) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1184
  unfolding minus_ereal_def divide_ereal_def using assms by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1185
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1186
lemma borel_measurable_ereal_setsum[simp, intro]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1187
  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1188
  assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1189
  shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1190
proof cases
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1191
  assume "finite S"
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1192
  thus ?thesis using assms
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1193
    by induct auto
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1194
qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1195
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1196
lemma borel_measurable_ereal_setprod[simp, intro]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1197
  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1198
  assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1199
  shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1200
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1201
  assume "finite S"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1202
  thus ?thesis using assms by induct auto
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1203
qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1204
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1205
lemma borel_measurable_SUP[simp, intro]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1206
  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> ereal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1207
  assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  1208
  shows "(\<lambda>x. SUP i : A. f i x) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1209
  unfolding borel_measurable_ereal_iff_ge
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1210
proof
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1211
  fix a
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1212
  have "?sup -` {a<..} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46731
diff changeset
  1213
    by (auto simp: less_SUP_iff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1214
  then show "?sup -` {a<..} \<inter> space M \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1215
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1216
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1217
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1218
lemma borel_measurable_INF[simp, intro]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1219
  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> ereal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1220
  assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  1221
  shows "(\<lambda>x. INF i : A. f i x) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1222
  unfolding borel_measurable_ereal_iff_less
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1223
proof
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1224
  fix a
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1225
  have "?inf -` {..<a} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46731
diff changeset
  1226
    by (auto simp: INF_less_iff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1227
  then show "?inf -` {..<a} \<inter> space M \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1228
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1229
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1230
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1231
lemma
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1232
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1233
  assumes "\<And>i. f i \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1234
  shows borel_measurable_liminf[simp, intro]: "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1235
    and borel_measurable_limsup[simp, intro]: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1236
  unfolding liminf_SUPR_INFI limsup_INFI_SUPR using assms by auto
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1237
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1238
lemma borel_measurable_ereal_LIMSEQ:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1239
  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1240
  assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1241
  and u: "\<And>i. u i \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1242
  shows "u' \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1243
proof -
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1244
  have "\<And>x. x \<in> space M \<Longrightarrow> u' x = liminf (\<lambda>n. u n x)"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1245
    using u' by (simp add: lim_imp_Liminf[symmetric])
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1246
  then show ?thesis by (simp add: u cong: measurable_cong)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1247
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1248
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1249
lemma borel_measurable_psuminf[simp, intro]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1250
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1251
  assumes "\<And>i. f i \<in> borel_measurable M" and pos: "\<And>i x. x \<in> space M \<Longrightarrow> 0 \<le> f i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1252
  shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1253
  apply (subst measurable_cong)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1254
  apply (subst suminf_ereal_eq_SUPR)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1255
  apply (rule pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1256
  using assms by auto
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1257
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1258
section "LIMSEQ is borel measurable"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1259
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1260
lemma borel_measurable_LIMSEQ:
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1261
  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1262
  assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1263
  and u: "\<And>i. u i \<in> borel_measurable M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1264
  shows "u' \<in> borel_measurable M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1265
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1266
  have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45288
diff changeset
  1267
    using u' by (simp add: lim_imp_Liminf)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1268
  moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1269
    by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1270
  ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1271
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1272
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1273
lemma sets_Collect_Cauchy: 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1274
  fixes f :: "nat \<Rightarrow> 'a => real"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1275
  assumes f: "\<And>i. f i \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1276
  shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1277
  unfolding Cauchy_iff2 using f by (auto intro!: sets_Collect)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1278
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1279
lemma borel_measurable_lim:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1280
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1281
  assumes f: "\<And>i. f i \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1282
  shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1283
proof -
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1284
  have *: "\<And>x. lim (\<lambda>i. f i x) =
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1285
    (if Cauchy (\<lambda>i. f i x) then lim (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0) else (THE x. False))"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1286
    by (auto simp: lim_def convergent_eq_cauchy[symmetric])
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1287
  { fix x have "convergent (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1288
      by (cases "Cauchy (\<lambda>i. f i x)")
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1289
         (auto simp add: convergent_eq_cauchy[symmetric] convergent_def) }
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1290
  note convergent = this
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1291
  show ?thesis
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1292
    unfolding *
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1293
    apply (intro measurable_If sets_Collect_Cauchy f borel_measurable_const)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1294
    apply (rule borel_measurable_LIMSEQ)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1295
    apply (rule convergent_LIMSEQ_iff[THEN iffD1, OF convergent])
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1296
    apply (intro measurable_If sets_Collect_Cauchy f borel_measurable_const)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1297
    done
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1298
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1299
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1300
lemma borel_measurable_suminf:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1301
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1302
  assumes f: "\<And>i. f i \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1303
  shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1304
  unfolding suminf_def sums_def[abs_def] lim_def[symmetric]
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1305
  by (simp add: f borel_measurable_lim)
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1306
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1307
end