src/HOL/Probability/Borel_Space.thy
author hoelzl
Thu, 13 Nov 2014 17:19:52 +0100
changeset 59000 6eb0725503fc
parent 58876 1888e3cb8048
child 59088 ff2bd4a14ddb
permissions -rw-r--r--
import general theorems from AFP/Markov_Models
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
58876
1888e3cb8048 modernized header;
wenzelm
parents: 58656
diff changeset
     6
section {*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
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents: 50245
diff changeset
     9
imports
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents: 50245
diff changeset
    10
  Measurable
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents: 50245
diff changeset
    11
  "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    12
begin
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    13
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56993
diff changeset
    14
subsection {* Generic Borel spaces *}
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    15
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    16
definition borel :: "'a::topological_space measure" where
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    17
  "borel = sigma UNIV {S. open S}"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    18
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    19
abbreviation "borel_measurable M \<equiv> measurable M borel"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    20
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    21
lemma in_borel_measurable:
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    22
   "f \<in> borel_measurable M \<longleftrightarrow>
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    23
    (\<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
    24
  by (auto simp add: measurable_def borel_def)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    25
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    26
lemma in_borel_measurable_borel:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    27
   "f \<in> borel_measurable M \<longleftrightarrow>
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    28
    (\<forall>S \<in> sets borel.
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    29
      f -` S \<inter> space M \<in> sets M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    30
  by (auto simp add: measurable_def borel_def)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    31
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    32
lemma space_borel[simp]: "space borel = UNIV"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    33
  unfolding borel_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    34
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
    35
lemma space_in_borel[measurable]: "UNIV \<in> sets borel"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
    36
  unfolding borel_def by auto
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
    37
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
    38
lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
    39
  unfolding borel_def by (rule sets_measure_of) simp
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
    40
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents: 50245
diff changeset
    41
lemma pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \<Longrightarrow> {x. P x} \<in> sets borel"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
    42
  unfolding borel_def pred_def by auto
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
    43
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    44
lemma borel_open[measurable (raw generic)]:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    45
  assumes "open A" shows "A \<in> sets borel"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    46
proof -
44537
c10485a6a7af make HOL-Probability respect set/pred distinction
huffman
parents: 44282
diff changeset
    47
  have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms .
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    48
  thus ?thesis unfolding borel_def by auto
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    49
qed
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    50
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    51
lemma borel_closed[measurable (raw generic)]:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    52
  assumes "closed A" shows "A \<in> sets borel"
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    53
proof -
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    54
  have "space borel - (- A) \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    55
    using assms unfolding closed_def by (blast intro: borel_open)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    56
  thus ?thesis by simp
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    57
qed
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    58
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    59
lemma borel_singleton[measurable]:
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    60
  "A \<in> sets borel \<Longrightarrow> insert x A \<in> sets (borel :: 'a::t1_space measure)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
    61
  unfolding insert_def by (rule sets.Un) auto
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
    62
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    63
lemma borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
    64
  unfolding Compl_eq_Diff_UNIV by simp
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
    65
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    66
lemma borel_measurable_vimage:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    67
  fixes f :: "'a \<Rightarrow> 'x::t2_space"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
    68
  assumes borel[measurable]: "f \<in> borel_measurable M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    69
  shows "f -` {x} \<inter> space M \<in> sets M"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
    70
  by simp
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    71
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    72
lemma borel_measurableI:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    73
  fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    74
  assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    75
  shows "f \<in> borel_measurable M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    76
  unfolding borel_def
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    77
proof (rule measurable_measure_of, simp_all)
44537
c10485a6a7af make HOL-Probability respect set/pred distinction
huffman
parents: 44282
diff changeset
    78
  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
    79
    using assms[of S] by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    80
qed
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    81
50021
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
    82
lemma borel_measurable_const:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    83
  "(\<lambda>x. c) \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    84
  by auto
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    85
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    86
lemma borel_measurable_indicator:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    87
  assumes A: "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
    88
  shows "indicator A \<in> borel_measurable M"
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46884
diff changeset
    89
  unfolding indicator_def [abs_def] using A
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    90
  by (auto intro!: measurable_If_set)
33533
40b44cb20c8c New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff changeset
    91
50096
7c9c5b1b6cd7 more measurability rules
hoelzl
parents: 50094
diff changeset
    92
lemma borel_measurable_count_space[measurable (raw)]:
7c9c5b1b6cd7 more measurability rules
hoelzl
parents: 50094
diff changeset
    93
  "f \<in> borel_measurable (count_space S)"
7c9c5b1b6cd7 more measurability rules
hoelzl
parents: 50094
diff changeset
    94
  unfolding measurable_def by auto
7c9c5b1b6cd7 more measurability rules
hoelzl
parents: 50094
diff changeset
    95
7c9c5b1b6cd7 more measurability rules
hoelzl
parents: 50094
diff changeset
    96
lemma borel_measurable_indicator'[measurable (raw)]:
7c9c5b1b6cd7 more measurability rules
hoelzl
parents: 50094
diff changeset
    97
  assumes [measurable]: "{x\<in>space M. f x \<in> A x} \<in> sets M"
7c9c5b1b6cd7 more measurability rules
hoelzl
parents: 50094
diff changeset
    98
  shows "(\<lambda>x. indicator (A x) (f x)) \<in> borel_measurable M"
50001
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49774
diff changeset
    99
  unfolding indicator_def[abs_def]
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49774
diff changeset
   100
  by (auto intro!: measurable_If)
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49774
diff changeset
   101
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   102
lemma borel_measurable_indicator_iff:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   103
  "(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
   104
    (is "?I \<in> borel_measurable M \<longleftrightarrow> _")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   105
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   106
  assume "?I \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   107
  then have "?I -` {1} \<inter> space M \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   108
    unfolding measurable_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   109
  also have "?I -` {1} \<inter> space M = A \<inter> space M"
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46884
diff changeset
   110
    unfolding indicator_def [abs_def] by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   111
  finally show "A \<inter> space M \<in> sets M" .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   112
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   113
  assume "A \<inter> space M \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   114
  moreover have "?I \<in> borel_measurable M \<longleftrightarrow>
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   115
    (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
   116
    by (intro measurable_cong) (auto simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   117
  ultimately show "?I \<in> borel_measurable M" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   118
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   119
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   120
lemma borel_measurable_subalgebra:
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41097
diff changeset
   121
  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
   122
  shows "f \<in> borel_measurable M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
   123
  using assms unfolding measurable_def by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
   124
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   125
lemma borel_measurable_restrict_space_iff_ereal:
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   126
  fixes f :: "'a \<Rightarrow> ereal"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   127
  assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   128
  shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   129
    (\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M"
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   130
  by (subst measurable_restrict_space_iff)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   131
     (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_cong)
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   132
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   133
lemma borel_measurable_restrict_space_iff:
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   134
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   135
  assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   136
  shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   137
    (\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> borel_measurable M"
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   138
  by (subst measurable_restrict_space_iff)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57447
diff changeset
   139
     (auto simp: indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] ac_simps cong del: if_cong)
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   140
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   141
lemma cbox_borel[measurable]: "cbox a b \<in> sets borel"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   142
  by (auto intro: borel_closed)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   143
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   144
lemma box_borel[measurable]: "box a b \<in> sets borel"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   145
  by (auto intro: borel_open)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   146
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   147
lemma borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   148
  by (auto intro: borel_closed dest!: compact_imp_closed)
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   149
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   150
lemma borel_measurable_continuous_on_if:
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   151
  assumes A: "A \<in> sets borel" and f: "continuous_on A f" and g: "continuous_on (- A) g"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   152
  shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> borel_measurable borel"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   153
proof (rule borel_measurableI)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   154
  fix S :: "'b set" assume "open S"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   155
  have "(\<lambda>x. if x \<in> A then f x else g x) -` S \<inter> space borel = (f -` S \<inter> A) \<union> (g -` S \<inter> -A)"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   156
    by (auto split: split_if_asm)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   157
  moreover obtain A' where "f -` S \<inter> A = A' \<inter> A" "open A'"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   158
    using continuous_on_open_invariant[THEN iffD1, rule_format, OF f `open S`] by metis
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   159
  moreover obtain B' where "g -` S \<inter> -A = B' \<inter> -A" "open B'"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   160
    using continuous_on_open_invariant[THEN iffD1, rule_format, OF g `open S`] by metis
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   161
  ultimately show "(\<lambda>x. if x \<in> A then f x else g x) -` S \<inter> space borel \<in> sets borel"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   162
    using A by auto
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   163
qed
f174712d0a84 better support for restrict_space
hoelzl
parents: 57036
diff changeset
   164
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   165
lemma borel_measurable_continuous_countable_exceptions:
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   166
  fixes f :: "'a::t1_space \<Rightarrow> 'b::topological_space"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   167
  assumes X: "countable X"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   168
  assumes "continuous_on (- X) f"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   169
  shows "f \<in> borel_measurable borel"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   170
proof (rule measurable_discrete_difference[OF _ X])
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   171
  have "X \<in> sets borel"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   172
    by (rule sets.countable[OF _ X]) auto
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   173
  then show "(\<lambda>x. if x \<in> X then undefined else f x) \<in> borel_measurable borel"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   174
    by (intro borel_measurable_continuous_on_if assms continuous_intros)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   175
qed auto
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   176
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   177
lemma borel_measurable_continuous_on1:
57138
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   178
  "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   179
  using borel_measurable_continuous_on_if[of UNIV f "\<lambda>_. undefined"]
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   180
  by (auto intro: continuous_on_const)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   181
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   182
lemma borel_measurable_continuous_on:
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   183
  assumes f: "continuous_on UNIV f" and g: "g \<in> borel_measurable M"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   184
  shows "(\<lambda>x. f (g x)) \<in> borel_measurable M"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   185
  using measurable_comp[OF g borel_measurable_continuous_on1[OF f]] by (simp add: comp_def)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   186
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   187
lemma borel_measurable_continuous_on_open':
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   188
  "continuous_on A f \<Longrightarrow> open A \<Longrightarrow>
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   189
    (\<lambda>x. if x \<in> A then f x else c) \<in> borel_measurable borel"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   190
  using borel_measurable_continuous_on_if[of A f "\<lambda>_. c"] by (auto intro: continuous_on_const)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   191
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   192
lemma borel_measurable_continuous_on_open:
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   193
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   194
  assumes cont: "continuous_on A f" "open A"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   195
  assumes g: "g \<in> borel_measurable M"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   196
  shows "(\<lambda>x. if g x \<in> A then f (g x) else c) \<in> borel_measurable M"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   197
  using measurable_comp[OF g borel_measurable_continuous_on_open'[OF cont], of c]
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   198
  by (simp add: comp_def)
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   199
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   200
lemma borel_measurable_continuous_on_indicator:
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   201
  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   202
  assumes A: "A \<in> sets borel" and f: "continuous_on A f"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   203
  shows "(\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable borel"
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   204
  using borel_measurable_continuous_on_if[OF assms, of "\<lambda>_. 0"]
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   205
  by (simp add: indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] continuous_on_const
7b3146180291 generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents: 57137
diff changeset
   206
           cong del: if_cong)
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   207
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   208
lemma borel_eq_countable_basis:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   209
  fixes B::"'a::topological_space set set"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   210
  assumes "countable B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   211
  assumes "topological_basis B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   212
  shows "borel = sigma UNIV B"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 50021
diff changeset
   213
  unfolding borel_def
635d73673b5e regularity of measures, therefore:
immler
parents: 50021
diff changeset
   214
proof (intro sigma_eqI sigma_sets_eqI, safe)
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   215
  interpret countable_basis using assms by unfold_locales
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   216
  fix X::"'a set" assume "open X"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   217
  from open_countable_basisE[OF this] guess B' . note B' = this
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   218
  then show "X \<in> sigma_sets UNIV B"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   219
    by (blast intro: sigma_sets_UNION `countable B` countable_subset)
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 50021
diff changeset
   220
next
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   221
  fix b assume "b \<in> B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   222
  hence "open b" by (rule topological_basis_open[OF assms(2)])
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   223
  thus "b \<in> sigma_sets UNIV (Collect open)" by auto
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 50021
diff changeset
   224
qed simp_all
635d73673b5e regularity of measures, therefore:
immler
parents: 50021
diff changeset
   225
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   226
lemma borel_measurable_Pair[measurable (raw)]:
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   227
  fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   228
  assumes f[measurable]: "f \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   229
  assumes g[measurable]: "g \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   230
  shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   231
proof (subst borel_eq_countable_basis)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   232
  let ?B = "SOME B::'b set set. countable B \<and> topological_basis B"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   233
  let ?C = "SOME B::'c set set. countable B \<and> topological_basis B"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   234
  let ?P = "(\<lambda>(b, c). b \<times> c) ` (?B \<times> ?C)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   235
  show "countable ?P" "topological_basis ?P"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   236
    by (auto intro!: countable_basis topological_basis_prod is_basis)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   237
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   238
  show "(\<lambda>x. (f x, g x)) \<in> measurable M (sigma UNIV ?P)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   239
  proof (rule measurable_measure_of)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   240
    fix S assume "S \<in> ?P"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   241
    then obtain b c where "b \<in> ?B" "c \<in> ?C" and S: "S = b \<times> c" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   242
    then have borel: "open b" "open c"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   243
      by (auto intro: is_basis topological_basis_open)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   244
    have "(\<lambda>x. (f x, g x)) -` S \<inter> space M = (f -` b \<inter> space M) \<inter> (g -` c \<inter> space M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   245
      unfolding S by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   246
    also have "\<dots> \<in> sets M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   247
      using borel by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   248
    finally show "(\<lambda>x. (f x, g x)) -` S \<inter> space M \<in> sets M" .
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   249
  qed auto
39087
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   250
qed
96984bf6fa5b Measurable on euclidean space is equiv. to measurable components
hoelzl
parents: 39083
diff changeset
   251
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   252
lemma borel_measurable_continuous_Pair:
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   253
  fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   254
  assumes [measurable]: "f \<in> borel_measurable M"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   255
  assumes [measurable]: "g \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   256
  assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   257
  shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   258
proof -
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   259
  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
   260
  show ?thesis
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   261
    unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   262
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   263
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56993
diff changeset
   264
subsection {* Borel spaces on euclidean spaces *}
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   265
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   266
lemma borel_measurable_inner[measurable (raw)]:
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   267
  fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_inner}"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   268
  assumes "f \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   269
  assumes "g \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   270
  shows "(\<lambda>x. f x \<bullet> g x) \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   271
  using assms
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56212
diff changeset
   272
  by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   273
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   274
lemma [measurable]:
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   275
  fixes a b :: "'a\<Colon>linorder_topology"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   276
  shows lessThan_borel: "{..< a} \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   277
    and greaterThan_borel: "{a <..} \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   278
    and greaterThanLessThan_borel: "{a<..<b} \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   279
    and atMost_borel: "{..a} \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   280
    and atLeast_borel: "{a..} \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   281
    and atLeastAtMost_borel: "{a..b} \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   282
    and greaterThanAtMost_borel: "{a<..b} \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   283
    and atLeastLessThan_borel: "{a..<b} \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   284
  unfolding greaterThanAtMost_def atLeastLessThan_def
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   285
  by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   286
                   closed_atMost closed_atLeast closed_atLeastAtMost)+
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   287
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   288
notation
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   289
  eucl_less (infix "<e" 50)
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   290
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   291
lemma box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}"
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   292
  and box_co: "{x. a \<le> x \<and> x <e b} = {a..} \<inter> {x. x <e b}"
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   293
  by auto
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   294
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   295
lemma eucl_ivals[measurable]:
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   296
  fixes a b :: "'a\<Colon>ordered_euclidean_space"
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   297
  shows "{x. x <e a} \<in> sets borel"
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   298
    and "{x. a <e x} \<in> sets borel"
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   299
    and "{..a} \<in> sets borel"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   300
    and "{a..} \<in> sets borel"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   301
    and "{a..b} \<in> sets borel"
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   302
    and  "{x. a <e x \<and> x \<le> b} \<in> sets borel"
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   303
    and "{x. a \<le> x \<and>  x <e b} \<in> sets borel"
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   304
  unfolding box_oc box_co
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   305
  by (auto intro: borel_open borel_closed)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   306
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   307
lemma open_Collect_less:
53216
ad2e09c30aa8 renamed inner_dense_linorder to dense_linorder
hoelzl
parents: 51683
diff changeset
   308
  fixes f g :: "'i::topological_space \<Rightarrow> 'a :: {dense_linorder, linorder_topology}"
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   309
  assumes "continuous_on UNIV f"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   310
  assumes "continuous_on UNIV g"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   311
  shows "open {x. f x < g x}"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   312
proof -
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   313
  have "open (\<Union>y. {x \<in> UNIV. f x \<in> {..< y}} \<inter> {x \<in> UNIV. g x \<in> {y <..}})" (is "open ?X")
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   314
    by (intro open_UN ballI open_Int continuous_open_preimage assms) auto
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   315
  also have "?X = {x. f x < g x}"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   316
    by (auto intro: dense)
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   317
  finally show ?thesis .
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   318
qed
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   319
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   320
lemma closed_Collect_le:
53216
ad2e09c30aa8 renamed inner_dense_linorder to dense_linorder
hoelzl
parents: 51683
diff changeset
   321
  fixes f g :: "'i::topological_space \<Rightarrow> 'a :: {dense_linorder, linorder_topology}"
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   322
  assumes f: "continuous_on UNIV f"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   323
  assumes g: "continuous_on UNIV g"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   324
  shows "closed {x. f x \<le> g x}"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   325
  using open_Collect_less[OF g f] unfolding not_less[symmetric] Collect_neg_eq open_closed .
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   326
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   327
lemma borel_measurable_less[measurable]:
53216
ad2e09c30aa8 renamed inner_dense_linorder to dense_linorder
hoelzl
parents: 51683
diff changeset
   328
  fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, dense_linorder, linorder_topology}"
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   329
  assumes "f \<in> borel_measurable M"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   330
  assumes "g \<in> borel_measurable M"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   331
  shows "{w \<in> space M. f w < g w} \<in> sets M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   332
proof -
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   333
  have "{w \<in> space M. f w < g w} = (\<lambda>x. (f x, g x)) -` {x. fst x < snd x} \<inter> space M"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   334
    by auto
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   335
  also have "\<dots> \<in> sets M"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   336
    by (intro measurable_sets[OF borel_measurable_Pair borel_open, OF assms open_Collect_less]
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56212
diff changeset
   337
              continuous_intros)
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   338
  finally show ?thesis .
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   339
qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   340
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   341
lemma
53216
ad2e09c30aa8 renamed inner_dense_linorder to dense_linorder
hoelzl
parents: 51683
diff changeset
   342
  fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, dense_linorder, linorder_topology}"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   343
  assumes f[measurable]: "f \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   344
  assumes g[measurable]: "g \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   345
  shows borel_measurable_le[measurable]: "{w \<in> space M. f w \<le> g w} \<in> sets M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   346
    and borel_measurable_eq[measurable]: "{w \<in> space M. f w = g w} \<in> sets M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   347
    and borel_measurable_neq: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   348
  unfolding eq_iff not_less[symmetric]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   349
  by measurable
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   350
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   351
lemma 
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   352
  fixes i :: "'a::{second_countable_topology, real_inner}"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   353
  shows hafspace_less_borel: "{x. a < x \<bullet> i} \<in> sets borel"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   354
    and hafspace_greater_borel: "{x. x \<bullet> i < a} \<in> sets borel"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   355
    and hafspace_less_eq_borel: "{x. a \<le> x \<bullet> i} \<in> sets borel"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   356
    and hafspace_greater_eq_borel: "{x. x \<bullet> i \<le> a} \<in> sets borel"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   357
  by simp_all
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   358
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   359
subsection "Borel space equals sigma algebras over intervals"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   360
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   361
lemma borel_sigma_sets_subset:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   362
  "A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   363
  using sets.sigma_sets_subset[of A borel] by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   364
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   365
lemma borel_eq_sigmaI1:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   366
  fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   367
  assumes borel_eq: "borel = sigma UNIV X"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   368
  assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (F ` A))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   369
  assumes F: "\<And>i. i \<in> A \<Longrightarrow> F i \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   370
  shows "borel = sigma UNIV (F ` A)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   371
  unfolding borel_def
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   372
proof (intro sigma_eqI antisym)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   373
  have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   374
    unfolding borel_def by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   375
  also have "\<dots> = sigma_sets UNIV X"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   376
    unfolding borel_eq by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   377
  also have "\<dots> \<subseteq> sigma_sets UNIV (F`A)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   378
    using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   379
  finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (F`A)" .
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   380
  show "sigma_sets UNIV (F`A) \<subseteq> sigma_sets UNIV {S. open S}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   381
    unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   382
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   383
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   384
lemma borel_eq_sigmaI2:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   385
  fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   386
    and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   387
  assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`B)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   388
  assumes X: "\<And>i j. (i, j) \<in> B \<Longrightarrow> G i j \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   389
  assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   390
  shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   391
  using assms
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   392
  by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` B" and F="(\<lambda>(i, j). F i j)"]) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   393
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   394
lemma borel_eq_sigmaI3:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   395
  fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   396
  assumes borel_eq: "borel = sigma UNIV X"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   397
  assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   398
  assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   399
  shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   400
  using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   401
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   402
lemma borel_eq_sigmaI4:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   403
  fixes F :: "'i \<Rightarrow> 'a::topological_space set"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   404
    and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   405
  assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`A)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   406
  assumes X: "\<And>i j. (i, j) \<in> A \<Longrightarrow> G i j \<in> sets (sigma UNIV (range F))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   407
  assumes F: "\<And>i. F i \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   408
  shows "borel = sigma UNIV (range F)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   409
  using assms by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` A" and F=F]) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   410
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   411
lemma borel_eq_sigmaI5:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   412
  fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   413
  assumes borel_eq: "borel = sigma UNIV (range G)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   414
  assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   415
  assumes F: "\<And>i j. F i j \<in> sets borel"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   416
  shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   417
  using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   418
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   419
lemma borel_eq_box:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   420
  "borel = sigma UNIV (range (\<lambda> (a, b). box a b :: 'a \<Colon> euclidean_space set))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   421
    (is "_ = ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   422
proof (rule borel_eq_sigmaI1[OF borel_def])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   423
  fix M :: "'a set" assume "M \<in> {S. open S}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   424
  then have "open M" by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   425
  show "M \<in> ?SIGMA"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   426
    apply (subst open_UNION_box[OF `open M`])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   427
    apply (safe intro!: sets.countable_UN' countable_PiE countable_Collect)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   428
    apply (auto intro: countable_rat)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   429
    done
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   430
qed (auto simp: box_def)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   431
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   432
lemma halfspace_gt_in_halfspace:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   433
  assumes i: "i \<in> A"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   434
  shows "{x\<Colon>'a. a < x \<bullet> i} \<in> 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   435
    sigma_sets UNIV ((\<lambda> (a, i). {x\<Colon>'a\<Colon>euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> A))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   436
  (is "?set \<in> ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   437
proof -
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   438
  interpret sigma_algebra UNIV ?SIGMA
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   439
    by (intro sigma_algebra_sigma_sets) simp_all
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   440
  have *: "?set = (\<Union>n. UNIV - {x\<Colon>'a. x \<bullet> i < a + 1 / real (Suc n)})"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   441
  proof (safe, simp_all add: not_less)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   442
    fix x :: 'a assume "a < x \<bullet> i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   443
    with reals_Archimedean[of "x \<bullet> i - a"]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   444
    obtain n where "a + 1 / real (Suc n) < x \<bullet> i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   445
      by (auto simp: inverse_eq_divide field_simps)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   446
    then show "\<exists>n. a + 1 / real (Suc n) \<le> x \<bullet> i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   447
      by (blast intro: less_imp_le)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   448
  next
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   449
    fix x n
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   450
    have "a < a + 1 / real (Suc n)" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   451
    also assume "\<dots> \<le> x"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   452
    finally show "a < x" .
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   453
  qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   454
  show "?set \<in> ?SIGMA" unfolding *
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   455
    by (auto del: Diff intro!: Diff i)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   456
qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   457
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   458
lemma borel_eq_halfspace_less:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   459
  "borel = sigma UNIV ((\<lambda>(a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> Basis))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   460
  (is "_ = ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   461
proof (rule borel_eq_sigmaI2[OF borel_eq_box])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   462
  fix a b :: 'a
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   463
  have "box a b = {x\<in>space ?SIGMA. \<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   464
    by (auto simp: box_def)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   465
  also have "\<dots> \<in> sets ?SIGMA"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   466
    by (intro sets.sets_Collect_conj sets.sets_Collect_finite_All sets.sets_Collect_const)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   467
       (auto intro!: halfspace_gt_in_halfspace countable_PiE countable_rat)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   468
  finally show "box a b \<in> sets ?SIGMA" .
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   469
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   470
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   471
lemma borel_eq_halfspace_le:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   472
  "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i \<le> a}) ` (UNIV \<times> Basis))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   473
  (is "_ = ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   474
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   475
  fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   476
  then have i: "i \<in> Basis" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   477
  have *: "{x::'a. x\<bullet>i < a} = (\<Union>n. {x. x\<bullet>i \<le> a - 1/real (Suc n)})"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   478
  proof (safe, simp_all)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   479
    fix x::'a assume *: "x\<bullet>i < a"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   480
    with reals_Archimedean[of "a - x\<bullet>i"]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   481
    obtain n where "x \<bullet> i < a - 1 / (real (Suc n))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   482
      by (auto simp: field_simps inverse_eq_divide)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   483
    then show "\<exists>n. x \<bullet> i \<le> a - 1 / (real (Suc n))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   484
      by (blast intro: less_imp_le)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   485
  next
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   486
    fix x::'a and n
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   487
    assume "x\<bullet>i \<le> a - 1 / real (Suc n)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   488
    also have "\<dots> < a" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   489
    finally show "x\<bullet>i < a" .
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   490
  qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   491
  show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding *
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   492
    by (safe intro!: sets.countable_UN) (auto intro: i)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   493
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   494
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   495
lemma borel_eq_halfspace_ge:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   496
  "borel = sigma UNIV ((\<lambda> (a, i). {x\<Colon>'a\<Colon>euclidean_space. a \<le> x \<bullet> i}) ` (UNIV \<times> Basis))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   497
  (is "_ = ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   498
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   499
  fix a :: real and i :: 'a assume i: "(a, i) \<in> UNIV \<times> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   500
  have *: "{x::'a. x\<bullet>i < a} = space ?SIGMA - {x::'a. a \<le> x\<bullet>i}" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   501
  show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding *
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   502
    using i by (safe intro!: sets.compl_sets) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   503
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   504
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   505
lemma borel_eq_halfspace_greater:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   506
  "borel = sigma UNIV ((\<lambda> (a, i). {x\<Colon>'a\<Colon>euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   507
  (is "_ = ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   508
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   509
  fix a :: real and i :: 'a assume "(a, i) \<in> (UNIV \<times> Basis)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   510
  then have i: "i \<in> Basis" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   511
  have *: "{x::'a. x\<bullet>i \<le> a} = space ?SIGMA - {x::'a. a < x\<bullet>i}" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   512
  show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding *
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   513
    by (safe intro!: sets.compl_sets) (auto intro: i)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   514
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   515
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   516
lemma borel_eq_atMost:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   517
  "borel = sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   518
  (is "_ = ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   519
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   520
  fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   521
  then have "i \<in> Basis" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   522
  then have *: "{x::'a. x\<bullet>i \<le> a} = (\<Union>k::nat. {.. (\<Sum>n\<in>Basis. (if n = i then a else real k)*\<^sub>R n)})"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   523
  proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   524
    fix x :: 'a
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   525
    from real_arch_simple[of "Max ((\<lambda>i. x\<bullet>i)`Basis)"] guess k::nat ..
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   526
    then have "\<And>i. i \<in> Basis \<Longrightarrow> x\<bullet>i \<le> real k"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   527
      by (subst (asm) Max_le_iff) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   528
    then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> x \<bullet> ia \<le> real k"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   529
      by (auto intro!: exI[of _ k])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   530
  qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   531
  show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding *
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   532
    by (safe intro!: sets.countable_UN) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   533
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   534
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   535
lemma borel_eq_greaterThan:
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   536
  "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {x. a <e x}))"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   537
  (is "_ = ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   538
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   539
  fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   540
  then have i: "i \<in> Basis" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   541
  have "{x::'a. x\<bullet>i \<le> a} = UNIV - {x::'a. a < x\<bullet>i}" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   542
  also have *: "{x::'a. a < x\<bullet>i} =
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   543
      (\<Union>k::nat. {x. (\<Sum>n\<in>Basis. (if n = i then a else -real k) *\<^sub>R n) <e x})" using i
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   544
  proof (safe, simp_all add: eucl_less_def split: split_if_asm)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   545
    fix x :: 'a
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   546
    from reals_Archimedean2[of "Max ((\<lambda>i. -x\<bullet>i)`Basis)"]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   547
    guess k::nat .. note k = this
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   548
    { fix i :: 'a assume "i \<in> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   549
      then have "-x\<bullet>i < real k"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   550
        using k by (subst (asm) Max_less_iff) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   551
      then have "- real k < x\<bullet>i" by simp }
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   552
    then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> -real k < x \<bullet> ia"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   553
      by (auto intro!: exI[of _ k])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   554
  qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   555
  finally show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   556
    apply (simp only:)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   557
    apply (safe intro!: sets.countable_UN sets.Diff)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   558
    apply (auto intro: sigma_sets_top)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   559
    done
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   560
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   561
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   562
lemma borel_eq_lessThan:
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   563
  "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {x. x <e a}))"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   564
  (is "_ = ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   565
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   566
  fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   567
  then have i: "i \<in> Basis" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   568
  have "{x::'a. a \<le> x\<bullet>i} = UNIV - {x::'a. x\<bullet>i < a}" by auto
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   569
  also have *: "{x::'a. x\<bullet>i < a} = (\<Union>k::nat. {x. x <e (\<Sum>n\<in>Basis. (if n = i then a else real k) *\<^sub>R n)})" using `i\<in> Basis`
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   570
  proof (safe, simp_all add: eucl_less_def split: split_if_asm)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   571
    fix x :: 'a
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   572
    from reals_Archimedean2[of "Max ((\<lambda>i. x\<bullet>i)`Basis)"]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   573
    guess k::nat .. note k = this
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   574
    { fix i :: 'a assume "i \<in> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   575
      then have "x\<bullet>i < real k"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   576
        using k by (subst (asm) Max_less_iff) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   577
      then have "x\<bullet>i < real k" by simp }
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   578
    then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> x \<bullet> ia < real k"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   579
      by (auto intro!: exI[of _ k])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   580
  qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   581
  finally show "{x. a \<le> x\<bullet>i} \<in> ?SIGMA"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   582
    apply (simp only:)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   583
    apply (safe intro!: sets.countable_UN sets.Diff)
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   584
    apply (auto intro: sigma_sets_top )
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   585
    done
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   586
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   587
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   588
lemma borel_eq_atLeastAtMost:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   589
  "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} \<Colon>'a\<Colon>ordered_euclidean_space set))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   590
  (is "_ = ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   591
proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   592
  fix a::'a
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   593
  have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   594
  proof (safe, simp_all add: eucl_le[where 'a='a])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   595
    fix x :: 'a
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   596
    from real_arch_simple[of "Max ((\<lambda>i. - x\<bullet>i)`Basis)"]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   597
    guess k::nat .. note k = this
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   598
    { fix i :: 'a assume "i \<in> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   599
      with k have "- x\<bullet>i \<le> real k"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   600
        by (subst (asm) Max_le_iff) (auto simp: field_simps)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   601
      then have "- real k \<le> x\<bullet>i" by simp }
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   602
    then show "\<exists>n::nat. \<forall>i\<in>Basis. - real n \<le> x \<bullet> i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   603
      by (auto intro!: exI[of _ k])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   604
  qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   605
  show "{..a} \<in> ?SIGMA" unfolding *
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   606
    by (safe intro!: sets.countable_UN)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   607
       (auto intro!: sigma_sets_top)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   608
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   609
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   610
lemma borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   611
proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   612
  fix i :: real
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   613
  have "{..i} = (\<Union>j::nat. {-j <.. i})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   614
    by (auto simp: minus_less_iff reals_Archimedean2)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   615
  also have "\<dots> \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   616
    by (intro sets.countable_nat_UN) auto 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   617
  finally show "{..i} \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))" .
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   618
qed simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
   619
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   620
lemma eucl_lessThan: "{x::real. x <e a} = lessThan a"
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   621
  by (simp add: eucl_less_def lessThan_def)
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   622
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   623
lemma borel_eq_atLeastLessThan:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   624
  "borel = sigma UNIV (range (\<lambda>(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA")
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   625
proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   626
  have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   627
  fix x :: real
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   628
  have "{..<x} = (\<Union>i::nat. {-real i ..< x})"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   629
    by (auto simp: move_uminus real_arch_simple)
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   630
  then show "{y. y <e x} \<in> ?SIGMA"
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
   631
    by (auto intro: sigma_sets.intros simp: eucl_lessThan)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   632
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   633
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   634
lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   635
  unfolding borel_def
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   636
proof (intro sigma_eqI sigma_sets_eqI, safe)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   637
  fix x :: "'a set" assume "open x"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   638
  hence "x = UNIV - (UNIV - x)" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   639
  also have "\<dots> \<in> sigma_sets UNIV (Collect closed)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   640
    by (rule sigma_sets.Compl)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   641
       (auto intro!: sigma_sets.Basic simp: `open x`)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   642
  finally show "x \<in> sigma_sets UNIV (Collect closed)" by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   643
next
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   644
  fix x :: "'a set" assume "closed x"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   645
  hence "x = UNIV - (UNIV - x)" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   646
  also have "\<dots> \<in> sigma_sets UNIV (Collect open)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   647
    by (rule sigma_sets.Compl)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   648
       (auto intro!: sigma_sets.Basic simp: `closed x`)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   649
  finally show "x \<in> sigma_sets UNIV (Collect open)" by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   650
qed simp_all
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   651
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   652
lemma borel_measurable_halfspacesI:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   653
  fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   654
  assumes F: "borel = sigma UNIV (F ` (UNIV \<times> Basis))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   655
  and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M" 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   656
  shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a::real. S a i \<in> sets M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   657
proof safe
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   658
  fix a :: real and i :: 'b assume i: "i \<in> Basis" and f: "f \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   659
  then show "S a i \<in> sets M" unfolding assms
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   660
    by (auto intro!: measurable_sets simp: assms(1))
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   661
next
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   662
  assume a: "\<forall>i\<in>Basis. \<forall>a. S a i \<in> sets M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   663
  then show "f \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   664
    by (auto intro!: measurable_measure_of simp: S_eq F)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   665
qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   666
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   667
lemma borel_measurable_iff_halfspace_le:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   668
  fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   669
  shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i \<le> a} \<in> sets M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   670
  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   671
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   672
lemma borel_measurable_iff_halfspace_less:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   673
  fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   674
  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i < a} \<in> sets M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   675
  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   676
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   677
lemma borel_measurable_iff_halfspace_ge:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   678
  fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   679
  shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a \<le> f w \<bullet> i} \<in> sets M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   680
  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   681
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   682
lemma borel_measurable_iff_halfspace_greater:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   683
  fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   684
  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a < f w \<bullet> i} \<in> sets M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   685
  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   686
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   687
lemma borel_measurable_iff_le:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   688
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   689
  using borel_measurable_iff_halfspace_le[where 'c=real] by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   690
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   691
lemma borel_measurable_iff_less:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   692
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   693
  using borel_measurable_iff_halfspace_less[where 'c=real] by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   694
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   695
lemma borel_measurable_iff_ge:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   696
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   697
  using borel_measurable_iff_halfspace_ge[where 'c=real]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   698
  by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   699
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   700
lemma borel_measurable_iff_greater:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   701
  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   702
  using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   703
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   704
lemma borel_measurable_euclidean_space:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   705
  fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   706
  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   707
proof safe
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   708
  assume f: "\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   709
  then show "f \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   710
    by (subst borel_measurable_iff_halfspace_le) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   711
qed auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   712
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   713
subsection "Borel measurable operators"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   714
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
   715
lemma borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
   716
  by (intro borel_measurable_continuous_on1 continuous_intros)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
   717
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   718
lemma borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   719
  by (rule borel_measurable_continuous_countable_exceptions[where X="{0}"])
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   720
     (auto intro!: continuous_on_sgn continuous_on_id)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   721
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   722
lemma borel_measurable_uminus[measurable (raw)]:
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   723
  fixes g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   724
  assumes g: "g \<in> borel_measurable M"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   725
  shows "(\<lambda>x. - g x) \<in> borel_measurable M"
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56212
diff changeset
   726
  by (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   727
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   728
lemma borel_measurable_add[measurable (raw)]:
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   729
  fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   730
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   731
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   732
  shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56212
diff changeset
   733
  using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   734
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   735
lemma borel_measurable_setsum[measurable (raw)]:
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   736
  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   737
  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
   738
  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
   739
proof cases
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   740
  assume "finite S"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   741
  thus ?thesis using assms by induct auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   742
qed simp
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   743
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   744
lemma borel_measurable_diff[measurable (raw)]:
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   745
  fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   746
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   747
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   748
  shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53216
diff changeset
   749
  using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   750
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   751
lemma borel_measurable_times[measurable (raw)]:
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   752
  fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   753
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   754
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   755
  shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56212
diff changeset
   756
  using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   757
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   758
lemma borel_measurable_setprod[measurable (raw)]:
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   759
  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, real_normed_field}"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   760
  assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   761
  shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   762
proof cases
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   763
  assume "finite S"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   764
  thus ?thesis using assms by induct auto
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   765
qed simp
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   766
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   767
lemma borel_measurable_dist[measurable (raw)]:
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   768
  fixes g f :: "'a \<Rightarrow> 'b::{second_countable_topology, metric_space}"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   769
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   770
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   771
  shows "(\<lambda>x. dist (f x) (g x)) \<in> borel_measurable M"
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56212
diff changeset
   772
  using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   773
  
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   774
lemma borel_measurable_scaleR[measurable (raw)]:
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   775
  fixes g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   776
  assumes f: "f \<in> borel_measurable M"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   777
  assumes g: "g \<in> borel_measurable M"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   778
  shows "(\<lambda>x. f x *\<^sub>R g x) \<in> borel_measurable M"
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56212
diff changeset
   779
  using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   780
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   781
lemma affine_borel_measurable_vector:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   782
  fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   783
  assumes "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   784
  shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   785
proof (rule borel_measurableI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   786
  fix S :: "'x set" assume "open S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   787
  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
   788
  proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   789
    assume "b \<noteq> 0"
44537
c10485a6a7af make HOL-Probability respect set/pred distinction
huffman
parents: 44282
diff changeset
   790
    with `open S` have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S")
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53216
diff changeset
   791
      using open_affinity [of S "inverse b" "- a /\<^sub>R b"]
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53216
diff changeset
   792
      by (auto simp: algebra_simps)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   793
    hence "?S \<in> sets borel" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   794
    moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   795
    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
   796
      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
   797
    ultimately show ?thesis using assms unfolding in_borel_measurable_borel
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   798
      by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   799
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   800
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   801
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   802
lemma borel_measurable_const_scaleR[measurable (raw)]:
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   803
  "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. b *\<^sub>R f x ::'a::real_normed_vector) \<in> borel_measurable M"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   804
  using affine_borel_measurable_vector[of f M 0 b] by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   805
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   806
lemma borel_measurable_const_add[measurable (raw)]:
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   807
  "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. a + f x ::'a::real_normed_vector) \<in> borel_measurable M"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   808
  using affine_borel_measurable_vector[of f M a 1] by simp
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   809
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   810
lemma borel_measurable_inverse[measurable (raw)]:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   811
  fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   812
  assumes f: "f \<in> borel_measurable M"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   813
  shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   814
  apply (rule measurable_compose[OF f])
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   815
  apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"])
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   816
  apply (auto intro!: continuous_on_inverse continuous_on_id)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   817
  done
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   818
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   819
lemma borel_measurable_divide[measurable (raw)]:
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   820
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   821
    (\<lambda>x. f x / g x::'b::{second_countable_topology, real_normed_div_algebra}) \<in> borel_measurable M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   822
  by (simp add: divide_inverse)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   823
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   824
lemma borel_measurable_max[measurable (raw)]:
53216
ad2e09c30aa8 renamed inner_dense_linorder to dense_linorder
hoelzl
parents: 51683
diff changeset
   825
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. max (g x) (f x) :: 'b::{second_countable_topology, dense_linorder, linorder_topology}) \<in> borel_measurable M"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   826
  by (simp add: max_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   827
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   828
lemma borel_measurable_min[measurable (raw)]:
53216
ad2e09c30aa8 renamed inner_dense_linorder to dense_linorder
hoelzl
parents: 51683
diff changeset
   829
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. min (g x) (f x) :: 'b::{second_countable_topology, dense_linorder, linorder_topology}) \<in> borel_measurable M"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   830
  by (simp add: min_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   831
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   832
lemma borel_measurable_Min[measurable (raw)]:
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   833
  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Min ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, dense_linorder, linorder_topology}) \<in> borel_measurable M"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   834
proof (induct I rule: finite_induct)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   835
  case (insert i I) then show ?case
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   836
    by (cases "I = {}") auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   837
qed auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   838
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   839
lemma borel_measurable_Max[measurable (raw)]:
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   840
  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Max ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, dense_linorder, linorder_topology}) \<in> borel_measurable M"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   841
proof (induct I rule: finite_induct)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   842
  case (insert i I) then show ?case
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   843
    by (cases "I = {}") auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   844
qed auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   845
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   846
lemma borel_measurable_abs[measurable (raw)]:
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   847
  "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   848
  unfolding abs_real_def by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
   849
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   850
lemma borel_measurable_nth[measurable (raw)]:
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   851
  "(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50419
diff changeset
   852
  by (simp add: cart_eq_inner_axis)
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41025
diff changeset
   853
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   854
lemma convex_measurable:
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   855
  fixes A :: "'a :: ordered_euclidean_space set"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   856
  assumes X: "X \<in> borel_measurable M" "X ` space M \<subseteq> A" "open A"
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   857
  assumes q: "convex_on A q"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   858
  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
   859
proof -
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   860
  have "(\<lambda>x. if X x \<in> A then q (X x) else 0) \<in> borel_measurable M" (is "?qX")
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   861
  proof (rule borel_measurable_continuous_on_open[OF _ _ X(1)])
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   862
    show "open A" by fact
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
   863
    from this q show "continuous_on A q"
42990
3706951a6421 composition of convex and measurable function is measurable
hoelzl
parents: 42950
diff changeset
   864
      by (rule convex_on_continuous)
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   865
  qed
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   866
  also 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
   867
    using X by (intro measurable_cong) auto
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   868
  finally show ?thesis .
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   869
qed
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   870
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   871
lemma borel_measurable_ln[measurable (raw)]:
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   872
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   873
  shows "(\<lambda>x. ln (f x)) \<in> borel_measurable M"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   874
  apply (rule measurable_compose[OF f])
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   875
  apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"])
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   876
  apply (auto intro!: continuous_on_ln continuous_on_id)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57259
diff changeset
   877
  done
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   878
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   879
lemma borel_measurable_log[measurable (raw)]:
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   880
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log (g x) (f x)) \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   881
  unfolding log_def by auto
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
   882
58656
7f14d5d9b933 relaxed class constraints for exp
immler
parents: 57514
diff changeset
   883
lemma borel_measurable_exp[measurable]:
7f14d5d9b933 relaxed class constraints for exp
immler
parents: 57514
diff changeset
   884
  "(exp::'a::{real_normed_field,banach}\<Rightarrow>'a) \<in> borel_measurable borel"
51478
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51351
diff changeset
   885
  by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_exp)
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50387
diff changeset
   886
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   887
lemma measurable_real_floor[measurable]:
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   888
  "(floor :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"
47761
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   889
proof -
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   890
  have "\<And>a x. \<lfloor>x\<rfloor> = a \<longleftrightarrow> (real a \<le> x \<and> x < real (a + 1))"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   891
    by (auto intro: floor_eq2)
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   892
  then show ?thesis
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   893
    by (auto simp: vimage_def measurable_count_space_eq2_countable)
47761
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   894
qed
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   895
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   896
lemma measurable_real_natfloor[measurable]:
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   897
  "(natfloor :: real \<Rightarrow> nat) \<in> measurable borel (count_space UNIV)"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   898
  by (simp add: natfloor_def[abs_def])
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   899
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   900
lemma measurable_real_ceiling[measurable]:
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   901
  "(ceiling :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   902
  unfolding ceiling_def[abs_def] by simp
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   903
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   904
lemma borel_measurable_real_floor: "(\<lambda>x::real. real \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   905
  by simp
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   906
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   907
lemma borel_measurable_real_natfloor:
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   908
  "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. real (natfloor (f x))) \<in> borel_measurable M"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   909
  by simp
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   910
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   911
lemma borel_measurable_root [measurable]: "(root n) \<in> borel_measurable borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   912
  by (intro borel_measurable_continuous_on1 continuous_intros)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   913
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   914
lemma borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   915
  by (intro borel_measurable_continuous_on1 continuous_intros)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   916
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   917
lemma borel_measurable_power [measurable (raw)]:
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   918
   fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   919
   assumes f: "f \<in> borel_measurable M"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   920
   shows "(\<lambda>x. (f x) ^ n) \<in> borel_measurable M"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   921
   by (intro borel_measurable_continuous_on [OF _ f] continuous_intros)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   922
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   923
lemma borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   924
  by (intro borel_measurable_continuous_on1 continuous_intros)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   925
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   926
lemma borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   927
  by (intro borel_measurable_continuous_on1 continuous_intros)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   928
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   929
lemma borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   930
  by (intro borel_measurable_continuous_on1 continuous_intros)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   931
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   932
lemma borel_measurable_sin [measurable]: "sin \<in> borel_measurable borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   933
  by (intro borel_measurable_continuous_on1 continuous_intros)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   934
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   935
lemma borel_measurable_cos [measurable]: "cos \<in> borel_measurable borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   936
  by (intro borel_measurable_continuous_on1 continuous_intros)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   937
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   938
lemma borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   939
  by (intro borel_measurable_continuous_on1 continuous_intros)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57138
diff changeset
   940
57259
3a448982a74a add more derivative and continuity rules for complex-values functions
hoelzl
parents: 57235
diff changeset
   941
lemma borel_measurable_complex_iff:
3a448982a74a add more derivative and continuity rules for complex-values functions
hoelzl
parents: 57235
diff changeset
   942
  "f \<in> borel_measurable M \<longleftrightarrow>
3a448982a74a add more derivative and continuity rules for complex-values functions
hoelzl
parents: 57235
diff changeset
   943
    (\<lambda>x. Re (f x)) \<in> borel_measurable M \<and> (\<lambda>x. Im (f x)) \<in> borel_measurable M"
3a448982a74a add more derivative and continuity rules for complex-values functions
hoelzl
parents: 57235
diff changeset
   944
  apply auto
3a448982a74a add more derivative and continuity rules for complex-values functions
hoelzl
parents: 57235
diff changeset
   945
  apply (subst fun_complex_eq)
3a448982a74a add more derivative and continuity rules for complex-values functions
hoelzl
parents: 57235
diff changeset
   946
  apply (intro borel_measurable_add)
3a448982a74a add more derivative and continuity rules for complex-values functions
hoelzl
parents: 57235
diff changeset
   947
  apply auto
3a448982a74a add more derivative and continuity rules for complex-values functions
hoelzl
parents: 57235
diff changeset
   948
  done
3a448982a74a add more derivative and continuity rules for complex-values functions
hoelzl
parents: 57235
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
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   952
lemma borel_measurable_ereal[measurable (raw)]:
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
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   956
lemma borel_measurable_real_of_ereal[measurable (raw)]:
49774
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 -
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   975
  let ?F = "\<lambda>x. if f x = \<infinity> then H \<infinity> else if f x = - \<infinity> then H (-\<infinity>) else H (ereal (real (f x)))"
49774
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 }
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   977
  with f H show ?thesis by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   978
qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
   979
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   980
lemma
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   981
  fixes f :: "'a \<Rightarrow> ereal" assumes f[measurable]: "f \<in> borel_measurable M"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   982
  shows borel_measurable_ereal_abs[measurable(raw)]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   983
    and borel_measurable_ereal_inverse[measurable(raw)]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   984
    and borel_measurable_uminus_ereal[measurable(raw)]: "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   985
  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
   986
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   987
lemma borel_measurable_uminus_eq_ereal[simp]:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   988
  "(\<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
   989
proof
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   990
  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
   991
qed auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   992
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   993
lemma set_Collect_ereal2:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   994
  fixes f g :: "'a \<Rightarrow> ereal" 
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   995
  assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   996
  assumes g: "g \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
   997
  assumes H: "{x \<in> space M. H (ereal (real (f x))) (ereal (real (g x)))} \<in> sets M"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   998
    "{x \<in> space borel. H (-\<infinity>) (ereal x)} \<in> sets borel"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   999
    "{x \<in> space borel. H (\<infinity>) (ereal x)} \<in> sets borel"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1000
    "{x \<in> space borel. H (ereal x) (-\<infinity>)} \<in> sets borel"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1001
    "{x \<in> space borel. H (ereal x) (\<infinity>)} \<in> sets borel"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1002
  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
  1003
proof -
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1004
  let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = -\<infinity> then H y (-\<infinity>) else H y (ereal (real (g x)))"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1005
  let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = -\<infinity> then ?G (-\<infinity>) x else ?G (ereal (real (f x))) x"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1006
  { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1007
  note * = this
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1008
  from assms show ?thesis
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1009
    by (subst *) (simp del: space_borel split del: split_if)
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1010
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1011
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1012
lemma borel_measurable_ereal_iff:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1013
  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
  1014
proof
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1015
  assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1016
  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
  1017
  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
  1018
qed auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1019
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1020
lemma borel_measurable_ereal_iff_real:
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1021
  fixes f :: "'a \<Rightarrow> ereal"
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1022
  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
  1023
    ((\<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
  1024
proof safe
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1025
  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
  1026
  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
  1027
  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
  1028
  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
  1029
  have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1030
  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
  1031
  finally show "f \<in> borel_measurable M" .
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1032
qed simp_all
41830
719b0a517c33 log is borel measurable
hoelzl
parents: 41545
diff changeset
  1033
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1034
lemma borel_measurable_eq_atMost_ereal:
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1035
  fixes f :: "'a \<Rightarrow> ereal"
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1036
  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
  1037
proof (intro iffI allI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1038
  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
  1039
  show "f \<in> borel_measurable M"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1040
    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
  1041
  proof (intro conjI allI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1042
    fix a :: real
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1043
    { 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
  1044
      have "x = \<infinity>"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1045
      proof (rule ereal_top)
44666
8670a39d4420 remove more duplicate lemmas
huffman
parents: 44537
diff changeset
  1046
        fix B from reals_Archimedean2[of B] guess n ..
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1047
        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
  1048
        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
  1049
      qed }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1050
    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
  1051
      by (auto simp: not_le)
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1052
    then show "f -` {\<infinity>} \<inter> space M \<in> sets M" using pos
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1053
      by (auto simp del: UN_simps)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1054
    moreover
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1055
    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
  1056
    then show "f -` {-\<infinity>} \<inter> space M \<in> sets M" using pos by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1057
    moreover have "{x\<in>space M. f x \<le> ereal a} \<in> sets M"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1058
      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
  1059
    moreover have "{w \<in> space M. real (f w) \<le> a} =
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1060
      (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
  1061
      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
  1062
      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
  1063
    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
  1064
  qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1065
qed (simp add: measurable_sets)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 35347
diff changeset
  1066
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1067
lemma borel_measurable_eq_atLeast_ereal:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1068
  "(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
  1069
proof
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1070
  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
  1071
  moreover have "\<And>a. (\<lambda>x. - f x) -` {..a} = f -` {-a ..}"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1072
    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
  1073
  ultimately have "(\<lambda>x. - f x) \<in> borel_measurable M"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1074
    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
  1075
  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
  1076
qed (simp add: measurable_sets)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 35347
diff changeset
  1077
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1078
lemma greater_eq_le_measurable:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1079
  fixes f :: "'a \<Rightarrow> 'c::linorder"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1080
  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
  1081
proof
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1082
  assume "f -` {a ..} \<inter> space M \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1083
  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
  1084
  ultimately show "f -` {..< a} \<inter> space M \<in> sets M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1085
next
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1086
  assume "f -` {..< a} \<inter> space M \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1087
  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
  1088
  ultimately show "f -` {a ..} \<inter> space M \<in> sets M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1089
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1090
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1091
lemma borel_measurable_ereal_iff_less:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1092
  "(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
  1093
  unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable ..
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1094
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1095
lemma less_eq_ge_measurable:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1096
  fixes f :: "'a \<Rightarrow> 'c::linorder"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1097
  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
  1098
proof
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1099
  assume "f -` {a <..} \<inter> space M \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1100
  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
  1101
  ultimately show "f -` {..a} \<inter> space M \<in> sets M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1102
next
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1103
  assume "f -` {..a} \<inter> space M \<in> sets M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1104
  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
  1105
  ultimately show "f -` {a <..} \<inter> space M \<in> sets M" by auto
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1106
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1107
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1108
lemma borel_measurable_ereal_iff_ge:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1109
  "(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
  1110
  unfolding borel_measurable_eq_atMost_ereal less_eq_ge_measurable ..
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1111
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1112
lemma borel_measurable_ereal2:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1113
  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
  1114
  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
  1115
  assumes g: "g \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1116
  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
  1117
    "(\<lambda>x. H (-\<infinity>) (ereal (real (g x)))) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1118
    "(\<lambda>x. H (\<infinity>) (ereal (real (g x)))) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1119
    "(\<lambda>x. H (ereal (real (f x))) (-\<infinity>)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1120
    "(\<lambda>x. H (ereal (real (f x))) (\<infinity>)) \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1121
  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
  1122
proof -
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1123
  let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = - \<infinity> then H y (-\<infinity>) else H y (ereal (real (g x)))"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1124
  let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = - \<infinity> then ?G (-\<infinity>) x else ?G (ereal (real (f x))) x"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1125
  { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1126
  note * = this
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1127
  from assms show ?thesis unfolding * by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1128
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1129
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1130
lemma
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1131
  fixes f :: "'a \<Rightarrow> ereal" assumes f: "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1132
  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
  1133
    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
  1134
  using f by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1135
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1136
lemma [measurable(raw)]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1137
  fixes f :: "'a \<Rightarrow> ereal"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1138
  assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1139
  shows borel_measurable_ereal_add: "(\<lambda>x. f x + g x) \<in> borel_measurable M"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1140
    and borel_measurable_ereal_times: "(\<lambda>x. f x * g x) \<in> borel_measurable M"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1141
    and borel_measurable_ereal_min: "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1142
    and borel_measurable_ereal_max: "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1143
  by (simp_all add: borel_measurable_ereal2 min_def max_def)
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1144
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1145
lemma [measurable(raw)]:
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1146
  fixes f g :: "'a \<Rightarrow> ereal"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1147
  assumes "f \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1148
  assumes "g \<in> borel_measurable M"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1149
  shows borel_measurable_ereal_diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1150
    and borel_measurable_ereal_divide: "(\<lambda>x. f x / g x) \<in> borel_measurable M"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1151
  using assms by (simp_all add: minus_ereal_def divide_ereal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1152
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1153
lemma borel_measurable_ereal_setsum[measurable (raw)]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1154
  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1155
  assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1156
  shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1157
proof cases
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1158
  assume "finite S"
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1159
  thus ?thesis using assms
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1160
    by induct auto
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1161
qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1162
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1163
lemma borel_measurable_ereal_setprod[measurable (raw)]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1164
  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1165
  assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1166
  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
  1167
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1168
  assume "finite S"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1169
  thus ?thesis using assms by induct auto
843c40bbc379 integral over setprod
hoelzl
parents: 41083
diff changeset
  1170
qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1171
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1172
lemma borel_measurable_SUP[measurable (raw)]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1173
  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> ereal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1174
  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
  1175
  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
  1176
  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
  1177
proof
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1178
  fix a
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1179
  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
  1180
    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
  1181
  then show "?sup -` {a<..} \<inter> space M \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1182
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1183
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1184
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1185
lemma borel_measurable_INF[measurable (raw)]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1186
  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> ereal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1187
  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
  1188
  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
  1189
  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
  1190
proof
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1191
  fix a
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1192
  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
  1193
    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
  1194
  then show "?inf -` {..<a} \<inter> space M \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1195
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1196
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37887
diff changeset
  1197
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1198
lemma [measurable (raw)]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1199
  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
  1200
  assumes "\<And>i. f i \<in> borel_measurable M"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1201
  shows borel_measurable_liminf: "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1202
    and borel_measurable_limsup: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
56212
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 54775
diff changeset
  1203
  unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1204
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50096
diff changeset
  1205
lemma sets_Collect_eventually_sequentially[measurable]:
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1206
  "(\<And>i. {x\<in>space M. P x i} \<in> sets M) \<Longrightarrow> {x\<in>space M. eventually (P x) sequentially} \<in> sets M"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1207
  unfolding eventually_sequentially by simp
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1208
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1209
lemma sets_Collect_ereal_convergent[measurable]: 
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1210
  fixes f :: "nat \<Rightarrow> 'a => ereal"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1211
  assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1212
  shows "{x\<in>space M. convergent (\<lambda>i. f i x)} \<in> sets M"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1213
  unfolding convergent_ereal by auto
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1214
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1215
lemma borel_measurable_extreal_lim[measurable (raw)]:
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1216
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1217
  assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1218
  shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1219
proof -
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1220
  have "\<And>x. lim (\<lambda>i. f i x) = (if convergent (\<lambda>i. f i x) then limsup (\<lambda>i. f i x) else (THE i. False))"
51351
dd1dd470690b generalized lemmas in Extended_Real_Limits
hoelzl
parents: 51106
diff changeset
  1221
    by (simp add: lim_def convergent_def convergent_limsup_cl)
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1222
  then show ?thesis
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1223
    by simp
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1224
qed
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1225
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1226
lemma borel_measurable_ereal_LIMSEQ:
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1227
  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1228
  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
  1229
  and u: "\<And>i. u i \<in> borel_measurable M"
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1230
  shows "u' \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1231
proof -
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1232
  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
  1233
    using u' by (simp add: lim_imp_Liminf[symmetric])
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1234
  with u show ?thesis by (simp cong: measurable_cong)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1235
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1236
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1237
lemma borel_measurable_extreal_suminf[measurable (raw)]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1238
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1239
  assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41969
diff changeset
  1240
  shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
  1241
  unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1242
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56993
diff changeset
  1243
subsection {* LIMSEQ is borel measurable *}
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1244
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1245
lemma borel_measurable_LIMSEQ:
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1246
  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1247
  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
  1248
  and u: "\<And>i. u i \<in> borel_measurable M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1249
  shows "u' \<in> borel_measurable M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1250
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1251
  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
  1252
    using u' by (simp add: lim_imp_Liminf)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1253
  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
  1254
    by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42990
diff changeset
  1255
  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
  1256
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39087
diff changeset
  1257
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1258
lemma borel_measurable_LIMSEQ_metric:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1259
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b :: metric_space"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1260
  assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1261
  assumes lim: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. f i x) ----> g x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1262
  shows "g \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1263
  unfolding borel_eq_closed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1264
proof (safe intro!: measurable_measure_of)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1265
  fix A :: "'b set" assume "closed A" 
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1266
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1267
  have [measurable]: "(\<lambda>x. infdist (g x) A) \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1268
  proof (rule borel_measurable_LIMSEQ)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1269
    show "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. infdist (f i x) A) ----> infdist (g x) A"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1270
      by (intro tendsto_infdist lim)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1271
    show "\<And>i. (\<lambda>x. infdist (f i x) A) \<in> borel_measurable M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1272
      by (intro borel_measurable_continuous_on[where f="\<lambda>x. infdist x A"]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1273
        continuous_at_imp_continuous_on ballI continuous_infdist isCont_ident) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1274
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1275
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1276
  show "g -` A \<inter> space M \<in> sets M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1277
  proof cases
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1278
    assume "A \<noteq> {}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1279
    then have "\<And>x. infdist x A = 0 \<longleftrightarrow> x \<in> A"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1280
      using `closed A` by (simp add: in_closed_iff_infdist_zero)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1281
    then have "g -` A \<inter> space M = {x\<in>space M. infdist (g x) A = 0}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1282
      by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1283
    also have "\<dots> \<in> sets M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1284
      by measurable
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1285
    finally show ?thesis .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1286
  qed simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1287
qed auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56371
diff changeset
  1288
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1289
lemma sets_Collect_Cauchy[measurable]: 
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 56994
diff changeset
  1290
  fixes f :: "nat \<Rightarrow> 'a => 'b::{metric_space, second_countable_topology}"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1291
  assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1292
  shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M"
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 56994
diff changeset
  1293
  unfolding metric_Cauchy_iff2 using f by auto
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1294
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1295
lemma borel_measurable_lim[measurable (raw)]:
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 56994
diff changeset
  1296
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1297
  assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1298
  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
  1299
proof -
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1300
  def u' \<equiv> "\<lambda>x. lim (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1301
  then have *: "\<And>x. lim (\<lambda>i. f i x) = (if Cauchy (\<lambda>i. f i x) then u' x else (THE x. False))"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1302
    by (auto simp: lim_def convergent_eq_cauchy[symmetric])
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1303
  have "u' \<in> borel_measurable M"
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 56994
diff changeset
  1304
  proof (rule borel_measurable_LIMSEQ_metric)
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1305
    fix x
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1306
    have "convergent (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1307
      by (cases "Cauchy (\<lambda>i. f i x)")
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1308
         (auto simp add: convergent_eq_cauchy[symmetric] convergent_def)
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1309
    then show "(\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0) ----> u' x"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1310
      unfolding u'_def 
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1311
      by (rule convergent_LIMSEQ_iff[THEN iffD1])
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1312
  qed measurable
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1313
  then show ?thesis
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1314
    unfolding * by measurable
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1315
qed
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1316
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1317
lemma borel_measurable_suminf[measurable (raw)]:
57036
22568fb89165 generalized Bochner integral over infinite sums
hoelzl
parents: 56994
diff changeset
  1318
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1319
  assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1320
  shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1321
  unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
49774
dfa8ddb874ce use continuity to show Borel-measurability
hoelzl
parents: 47761
diff changeset
  1322
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1323
lemma borel_measurable_sup[measurable (raw)]:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1324
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1325
    (\<lambda>x. sup (f x) (g x)::ereal) \<in> borel_measurable M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1326
  by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1327
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1328
lemma borel_measurable_lfp[consumes 1, case_names continuity step]:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1329
  fixes F :: "('a \<Rightarrow> ereal) \<Rightarrow> ('a \<Rightarrow> ereal)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1330
  assumes "Order_Continuity.continuous F"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1331
  assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1332
  shows "lfp F \<in> borel_measurable M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1333
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1334
  { fix i have "((F ^^ i) (\<lambda>t. bot)) \<in> borel_measurable M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1335
      by (induct i) (auto intro!: * simp: bot_fun_def) }
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1336
  then have "(\<lambda>x. SUP i. (F ^^ i) (\<lambda>t. bot) x) \<in> borel_measurable M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1337
    by measurable
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1338
  also have "(\<lambda>x. SUP i. (F ^^ i) (\<lambda>t. bot) x) = (SUP i. (F ^^ i) bot)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1339
    by (auto simp add: bot_fun_def)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1340
  also have "(SUP i. (F ^^ i) bot) = lfp F"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1341
    by (rule continuous_lfp[symmetric]) fact
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1342
  finally show ?thesis .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1343
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1344
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1345
(* Proof by Jeremy Avigad and Luke Serafin *)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1346
lemma isCont_borel:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1347
  fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1348
  shows "{x. isCont f x} \<in> sets borel"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1349
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1350
  let ?I = "\<lambda>j. inverse(real (Suc j))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1351
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1352
  { fix x
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1353
    have "isCont f x = (\<forall>i. \<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1354
      unfolding continuous_at_eps_delta
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1355
    proof safe
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1356
      fix i assume "\<forall>e>0. \<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1357
      moreover have "0 < ?I i / 2"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1358
        by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1359
      ultimately obtain d where d: "0 < d" "\<And>y. dist x y < d \<Longrightarrow> dist (f y) (f x) < ?I i / 2"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1360
        by (metis dist_commute)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1361
      then obtain j where j: "?I j < d"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1362
        by (metis reals_Archimedean)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1363
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1364
      show "\<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1365
      proof (safe intro!: exI[where x=j])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1366
        fix y z assume *: "dist x y < ?I j" "dist x z < ?I j"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1367
        have "dist (f y) (f z) \<le> dist (f y) (f x) + dist (f z) (f x)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1368
          by (rule dist_triangle2)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1369
        also have "\<dots> < ?I i / 2 + ?I i / 2"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1370
          by (intro add_strict_mono d less_trans[OF _ j] *)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1371
        also have "\<dots> \<le> ?I i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1372
          by (simp add: field_simps real_of_nat_Suc)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1373
        finally show "dist (f y) (f z) \<le> ?I i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1374
          by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1375
      qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1376
    next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1377
      fix e::real assume "0 < e"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1378
      then obtain n where n: "?I n < e"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1379
        by (metis reals_Archimedean)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1380
      assume "\<forall>i. \<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1381
      from this[THEN spec, of "Suc n"]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1382
      obtain j where j: "\<And>y z. dist x y < ?I j \<Longrightarrow> dist x z < ?I j \<Longrightarrow> dist (f y) (f z) \<le> ?I (Suc n)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1383
        by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1384
      
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1385
      show "\<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1386
      proof (safe intro!: exI[of _ "?I j"])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1387
        fix y assume "dist y x < ?I j"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1388
        then have "dist (f y) (f x) \<le> ?I (Suc n)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1389
          by (intro j) (auto simp: dist_commute)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1390
        also have "?I (Suc n) < ?I n"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1391
          by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1392
        also note n
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1393
        finally show "dist (f y) (f x) < e" .
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1394
      qed simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1395
    qed }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1396
  note * = this
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1397
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1398
  have **: "\<And>e y. open {x. dist x y < e}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1399
    using open_ball by (simp_all add: ball_def dist_commute)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1400
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1401
  have "{x\<in>space borel. isCont f x } \<in> sets borel"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1402
    unfolding *
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1403
    apply (intro sets.sets_Collect_countable_All sets.sets_Collect_countable_Ex)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1404
    apply (simp add: Collect_all_eq)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1405
    apply (intro borel_closed closed_INT ballI closed_Collect_imp open_Collect_conj **)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1406
    apply auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1407
    done
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1408
  then show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1409
    by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1410
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57275
diff changeset
  1411
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
  1412
no_notation
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
  1413
  eucl_less (infix "<e" 50)
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54230
diff changeset
  1414
51683
baefa3b461c2 generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
hoelzl
parents: 51478
diff changeset
  1415
end