| author | wenzelm | 
| Tue, 29 Aug 2023 16:42:08 +0200 | |
| changeset 78600 | afb83ba8d24c | 
| parent 78519 | f675e2a31682 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 63627 | 1  | 
(* Title: HOL/Analysis/Borel_Space.thy  | 
| 42067 | 2  | 
Author: Johannes Hölzl, TU München  | 
3  | 
Author: Armin Heller, TU München  | 
|
4  | 
*)  | 
|
| 38656 | 5  | 
|
| 
69722
 
b5163b2132c5
minor tagging updates in 13 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
6  | 
section \<open>Borel Space\<close>  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
7  | 
|
| 40859 | 8  | 
theory Borel_Space  | 
| 50387 | 9  | 
imports  | 
| 
63626
 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 
hoelzl 
parents: 
63566 
diff
changeset
 | 
10  | 
Measurable Derivative Ordered_Euclidean_Space Extended_Real_Limits  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
11  | 
begin  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
12  | 
|
| 
71025
 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 
immler 
parents: 
70688 
diff
changeset
 | 
13  | 
lemma is_interval_real_ereal_oo: "is_interval (real_of_ereal ` {N<..<M::ereal})"
 | 
| 
 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 
immler 
parents: 
70688 
diff
changeset
 | 
14  | 
by (auto simp: real_atLeastGreaterThan_eq)  | 
| 
 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 
immler 
parents: 
70688 
diff
changeset
 | 
15  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
16  | 
lemma sets_Collect_eventually_sequentially[measurable]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
17  | 
  "(\<And>i. {x\<in>space M. P x i} \<in> sets M) \<Longrightarrow> {x\<in>space M. eventually (P x) sequentially} \<in> sets M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
18  | 
unfolding eventually_sequentially by simp  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
19  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
20  | 
lemma topological_basis_trivial: "topological_basis {A. open A}"
 | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
21  | 
by (auto simp: topological_basis_def)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
22  | 
|
| 
69722
 
b5163b2132c5
minor tagging updates in 13 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
23  | 
proposition open_prod_generated: "open = generate_topology {A \<times> B | A B. open A \<and> open B}"
 | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
24  | 
proof -  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
25  | 
  have "{A \<times> B :: ('a \<times> 'b) set | A B. open A \<and> open B} = ((\<lambda>(a, b). a \<times> b) ` ({A. open A} \<times> {A. open A}))"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
26  | 
by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
27  | 
then show ?thesis  | 
| 62372 | 28  | 
by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis)  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
29  | 
qed  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
30  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
31  | 
proposition mono_on_imp_deriv_nonneg:  | 
| 
75607
 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 
desharna 
parents: 
74362 
diff
changeset
 | 
32  | 
assumes mono: "mono_on A f" and deriv: "(f has_real_derivative D) (at x)"  | 
| 62083 | 33  | 
assumes "x \<in> interior A"  | 
34  | 
shows "D \<ge> 0"  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
35  | 
proof (rule tendsto_lowerbound)  | 
| 62083 | 36  | 
let ?A' = "(\<lambda>y. y - x) ` interior A"  | 
37  | 
from deriv show "((\<lambda>h. (f (x + h) - f x) / h) \<longlongrightarrow> D) (at 0)"  | 
|
38  | 
by (simp add: field_has_derivative_at has_field_derivative_def)  | 
|
| 
75607
 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 
desharna 
parents: 
74362 
diff
changeset
 | 
39  | 
from mono have mono': "mono_on (interior A) f" by (rule mono_on_subset) (rule interior_subset)  | 
| 62083 | 40  | 
|
41  | 
show "eventually (\<lambda>h. (f (x + h) - f x) / h \<ge> 0) (at 0)"  | 
|
42  | 
proof (subst eventually_at_topological, intro exI conjI ballI impI)  | 
|
43  | 
have "open (interior A)" by simp  | 
|
| 67399 | 44  | 
hence "open ((+) (-x) ` interior A)" by (rule open_translation)  | 
45  | 
also have "((+) (-x) ` interior A) = ?A'" by auto  | 
|
| 62083 | 46  | 
finally show "open ?A'" .  | 
47  | 
next  | 
|
48  | 
from \<open>x \<in> interior A\<close> show "0 \<in> ?A'" by auto  | 
|
49  | 
next  | 
|
50  | 
fix h assume "h \<in> ?A'"  | 
|
51  | 
hence "x + h \<in> interior A" by auto  | 
|
52  | 
with mono' and \<open>x \<in> interior A\<close> show "(f (x + h) - f x) / h \<ge> 0"  | 
|
53  | 
by (cases h rule: linorder_cases[of _ 0])  | 
|
54  | 
(simp_all add: divide_nonpos_neg divide_nonneg_pos mono_onD field_simps)  | 
|
55  | 
qed  | 
|
56  | 
qed simp  | 
|
57  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
58  | 
proposition mono_on_ctble_discont:  | 
| 62083 | 59  | 
fixes f :: "real \<Rightarrow> real"  | 
60  | 
fixes A :: "real set"  | 
|
| 
75607
 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 
desharna 
parents: 
74362 
diff
changeset
 | 
61  | 
assumes "mono_on A f"  | 
| 62083 | 62  | 
  shows "countable {a\<in>A. \<not> continuous (at a within A) f}"
 | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
63  | 
proof -  | 
| 62083 | 64  | 
have mono: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"  | 
| 
75607
 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 
desharna 
parents: 
74362 
diff
changeset
 | 
65  | 
using \<open>mono_on A f\<close> by (simp add: mono_on_def)  | 
| 62083 | 66  | 
  have "\<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}. \<exists>q :: nat \<times> rat.
 | 
67  | 
(fst q = 0 \<and> of_rat (snd q) < f a \<and> (\<forall>x \<in> A. x < a \<longrightarrow> f x < of_rat (snd q))) \<or>  | 
|
68  | 
(fst q = 1 \<and> of_rat (snd q) > f a \<and> (\<forall>x \<in> A. x > a \<longrightarrow> f x > of_rat (snd q)))"  | 
|
69  | 
proof (clarsimp simp del: One_nat_def)  | 
|
70  | 
fix a assume "a \<in> A" assume "\<not> continuous (at a within A) f"  | 
|
71  | 
thus "\<exists>q1 q2.  | 
|
72  | 
q1 = 0 \<and> real_of_rat q2 < f a \<and> (\<forall>x\<in>A. x < a \<longrightarrow> f x < real_of_rat q2) \<or>  | 
|
73  | 
q1 = 1 \<and> f a < real_of_rat q2 \<and> (\<forall>x\<in>A. a < x \<longrightarrow> real_of_rat q2 < f x)"  | 
|
74  | 
proof (auto simp add: continuous_within order_tendsto_iff eventually_at)  | 
|
75  | 
fix l assume "l < f a"  | 
|
76  | 
then obtain q2 where q2: "l < of_rat q2" "of_rat q2 < f a"  | 
|
77  | 
using of_rat_dense by blast  | 
|
78  | 
assume * [rule_format]: "\<forall>d>0. \<exists>x\<in>A. x \<noteq> a \<and> dist x a < d \<and> \<not> l < f x"  | 
|
79  | 
from q2 have "real_of_rat q2 < f a \<and> (\<forall>x\<in>A. x < a \<longrightarrow> f x < real_of_rat q2)"  | 
|
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
80  | 
using q2 *[of "a - _"] dist_real_def mono by fastforce  | 
| 62083 | 81  | 
thus ?thesis by auto  | 
82  | 
next  | 
|
83  | 
fix u assume "u > f a"  | 
|
84  | 
then obtain q2 where q2: "f a < of_rat q2" "of_rat q2 < u"  | 
|
85  | 
using of_rat_dense by blast  | 
|
86  | 
assume *[rule_format]: "\<forall>d>0. \<exists>x\<in>A. x \<noteq> a \<and> dist x a < d \<and> \<not> u > f x"  | 
|
87  | 
from q2 have "real_of_rat q2 > f a \<and> (\<forall>x\<in>A. x > a \<longrightarrow> f x > real_of_rat q2)"  | 
|
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
88  | 
using q2 *[of "_ - a"] dist_real_def mono by fastforce  | 
| 62083 | 89  | 
thus ?thesis by auto  | 
90  | 
qed  | 
|
91  | 
qed  | 
|
| 74362 | 92  | 
  then obtain g :: "real \<Rightarrow> nat \<times> rat" where "\<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}.
 | 
| 62083 | 93  | 
(fst (g a) = 0 \<and> of_rat (snd (g a)) < f a \<and> (\<forall>x \<in> A. x < a \<longrightarrow> f x < of_rat (snd (g a)))) |  | 
94  | 
(fst (g a) = 1 \<and> of_rat (snd (g a)) > f a \<and> (\<forall>x \<in> A. x > a \<longrightarrow> f x > of_rat (snd (g a))))"  | 
|
| 74362 | 95  | 
by (rule bchoice [THEN exE]) blast  | 
| 62083 | 96  | 
hence g: "\<And>a x. a \<in> A \<Longrightarrow> \<not> continuous (at a within A) f \<Longrightarrow> x \<in> A \<Longrightarrow>  | 
97  | 
(fst (g a) = 0 \<and> of_rat (snd (g a)) < f a \<and> (x < a \<longrightarrow> f x < of_rat (snd (g a)))) |  | 
|
98  | 
(fst (g a) = 1 \<and> of_rat (snd (g a)) > f a \<and> (x > a \<longrightarrow> f x > of_rat (snd (g a))))"  | 
|
99  | 
by auto  | 
|
100  | 
  have "inj_on g {a\<in>A. \<not> continuous (at a within A) f}"
 | 
|
101  | 
proof (auto simp add: inj_on_def)  | 
|
102  | 
fix w z  | 
|
103  | 
assume 1: "w \<in> A" and 2: "\<not> continuous (at w within A) f" and  | 
|
104  | 
3: "z \<in> A" and 4: "\<not> continuous (at z within A) f" and  | 
|
105  | 
5: "g w = g z"  | 
|
| 62372 | 106  | 
from g [OF 1 2 3] g [OF 3 4 1] 5  | 
| 62083 | 107  | 
show "w = z" by auto  | 
108  | 
qed  | 
|
| 62372 | 109  | 
thus ?thesis  | 
110  | 
by (rule countableI')  | 
|
| 62083 | 111  | 
qed  | 
112  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
113  | 
lemma mono_on_ctble_discont_open:  | 
| 62083 | 114  | 
fixes f :: "real \<Rightarrow> real"  | 
115  | 
fixes A :: "real set"  | 
|
| 
75607
 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 
desharna 
parents: 
74362 
diff
changeset
 | 
116  | 
assumes "open A" "mono_on A f"  | 
| 62083 | 117  | 
  shows "countable {a\<in>A. \<not>isCont f a}"
 | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
118  | 
using continuous_within_open [OF _ \<open>open A\<close>] \<open>mono_on A f\<close>  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
119  | 
by (smt (verit, ccfv_threshold) Collect_cong mono_on_ctble_discont)  | 
| 62083 | 120  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
121  | 
lemma mono_ctble_discont:  | 
| 62083 | 122  | 
fixes f :: "real \<Rightarrow> real"  | 
123  | 
assumes "mono f"  | 
|
124  | 
  shows "countable {a. \<not> isCont f a}"
 | 
|
| 
75607
 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 
desharna 
parents: 
74362 
diff
changeset
 | 
125  | 
using assms mono_on_ctble_discont [of UNIV f] unfolding mono_on_def mono_def by auto  | 
| 62083 | 126  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
127  | 
lemma has_real_derivative_imp_continuous_on:  | 
| 62083 | 128  | 
assumes "\<And>x. x \<in> A \<Longrightarrow> (f has_real_derivative f' x) (at x)"  | 
129  | 
shows "continuous_on A f"  | 
|
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
130  | 
by (meson DERIV_isCont assms continuous_at_imp_continuous_on)  | 
| 62083 | 131  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
132  | 
lemma continuous_interval_vimage_Int:  | 
| 62083 | 133  | 
  assumes "continuous_on {a::real..b} g" and mono: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
 | 
134  | 
  assumes "a \<le> b" "(c::real) \<le> d" "{c..d} \<subseteq> {g a..g b}"
 | 
|
135  | 
  obtains c' d' where "{a..b} \<inter> g -` {c..d} = {c'..d'}" "c' \<le> d'" "g c' = c" "g d' = d"
 | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
136  | 
proof-  | 
| 63040 | 137  | 
  let ?A = "{a..b} \<inter> g -` {c..d}"
 | 
138  | 
from IVT'[of g a c b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)  | 
|
139  | 
obtain c'' where c'': "c'' \<in> ?A" "g c'' = c" by auto  | 
|
140  | 
from IVT'[of g a d b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)  | 
|
141  | 
obtain d'' where d'': "d'' \<in> ?A" "g d'' = d" by auto  | 
|
142  | 
  hence [simp]: "?A \<noteq> {}" by blast
 | 
|
| 62083 | 143  | 
|
| 63040 | 144  | 
define c' where "c' = Inf ?A"  | 
145  | 
define d' where "d' = Sup ?A"  | 
|
146  | 
  have "?A \<subseteq> {c'..d'}" unfolding c'_def d'_def
 | 
|
147  | 
by (intro subsetI) (auto intro: cInf_lower cSup_upper)  | 
|
148  | 
moreover from assms have "closed ?A"  | 
|
149  | 
    using continuous_on_closed_vimage[of "{a..b}" g] by (subst Int_commute) simp
 | 
|
150  | 
hence c'd'_in_set: "c' \<in> ?A" "d' \<in> ?A" unfolding c'_def d'_def  | 
|
151  | 
by ((intro closed_contains_Inf closed_contains_Sup, simp_all)[])+  | 
|
152  | 
  hence "{c'..d'} \<subseteq> ?A" using assms
 | 
|
153  | 
by (intro subsetI)  | 
|
154  | 
(auto intro!: order_trans[of c "g c'" "g x" for x] order_trans[of "g x" "g d'" d for x]  | 
|
155  | 
intro!: mono)  | 
|
156  | 
moreover have "c' \<le> d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto  | 
|
157  | 
moreover have "g c' \<le> c" "g d' \<ge> d"  | 
|
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
158  | 
using c'' d'' calculation by (metis IntE atLeastAtMost_iff mono order_class.order_eq_iff)+  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
159  | 
with c'd'_in_set have "g c' = c" "g d' = d"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
160  | 
by auto  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
161  | 
ultimately show ?thesis  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
162  | 
using that by blast  | 
| 62083 | 163  | 
qed  | 
164  | 
||
| 69683 | 165  | 
subsection \<open>Generic Borel spaces\<close>  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
166  | 
|
| 70136 | 167  | 
definition\<^marker>\<open>tag important\<close> (in topological_space) borel :: "'a measure" where  | 
| 47694 | 168  | 
  "borel = sigma UNIV {S. open S}"
 | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
169  | 
|
| 47694 | 170  | 
abbreviation "borel_measurable M \<equiv> measurable M borel"  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
171  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
172  | 
lemma in_borel_measurable:  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
173  | 
"f \<in> borel_measurable M \<longleftrightarrow>  | 
| 47694 | 174  | 
    (\<forall>S \<in> sigma_sets UNIV {S. open S}. f -` S \<inter> space M \<in> sets M)"
 | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
175  | 
by (auto simp add: measurable_def borel_def)  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
176  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
177  | 
lemma in_borel_measurable_borel:  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
178  | 
"f \<in> borel_measurable M \<longleftrightarrow> (\<forall>S \<in> sets borel. f -` S \<inter> space M \<in> sets M)"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
179  | 
by (auto simp add: measurable_def borel_def)  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
180  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
181  | 
lemma space_borel[simp]: "space borel = UNIV"  | 
| 40859 | 182  | 
unfolding borel_def by auto  | 
| 38656 | 183  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
184  | 
lemma space_in_borel[measurable]: "UNIV \<in> sets borel"  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
185  | 
unfolding borel_def by auto  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
186  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
187  | 
lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57138 
diff
changeset
 | 
188  | 
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
 | 
189  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
190  | 
lemma measurable_sets_borel:  | 
| 62083 | 191  | 
"\<lbrakk>f \<in> measurable borel M; A \<in> sets M\<rbrakk> \<Longrightarrow> f -` A \<in> sets borel"  | 
192  | 
by (drule (1) measurable_sets) simp  | 
|
193  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
194  | 
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
 | 
195  | 
unfolding borel_def pred_def by auto  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
196  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
197  | 
lemma borel_open[measurable (raw generic)]:  | 
| 40859 | 198  | 
assumes "open A" shows "A \<in> sets borel"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
199  | 
by (simp add: assms sets_borel)  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
200  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
201  | 
lemma borel_closed[measurable (raw generic)]:  | 
| 40859 | 202  | 
assumes "closed A" shows "A \<in> sets borel"  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
203  | 
proof -  | 
| 40859 | 204  | 
have "space borel - (- A) \<in> sets borel"  | 
205  | 
using assms unfolding closed_def by (blast intro: borel_open)  | 
|
| 38656 | 206  | 
thus ?thesis by simp  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
207  | 
qed  | 
| 
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
208  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
209  | 
lemma borel_singleton[measurable]:  | 
| 50003 | 210  | 
"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
 | 
211  | 
unfolding insert_def by (rule sets.Un) auto  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
212  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
213  | 
lemma sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV"
 | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
214  | 
by (simp add: set_eq_iff sets.countable)  | 
| 
64320
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64287 
diff
changeset
 | 
215  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
216  | 
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
 | 
217  | 
unfolding Compl_eq_Diff_UNIV by simp  | 
| 41830 | 218  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
219  | 
lemma borel_measurable_vimage:  | 
| 38656 | 220  | 
fixes f :: "'a \<Rightarrow> 'x::t2_space"  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
221  | 
assumes borel[measurable]: "f \<in> borel_measurable M"  | 
| 38656 | 222  | 
  shows "f -` {x} \<inter> space M \<in> sets M"
 | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
223  | 
by simp  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
224  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
225  | 
lemma borel_measurableI:  | 
| 61076 | 226  | 
fixes f :: "'a \<Rightarrow> 'x::topological_space"  | 
| 38656 | 227  | 
assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"  | 
228  | 
shows "f \<in> borel_measurable M"  | 
|
| 40859 | 229  | 
unfolding borel_def  | 
| 47694 | 230  | 
proof (rule measurable_measure_of, simp_all)  | 
| 
44537
 
c10485a6a7af
make HOL-Probability respect set/pred distinction
 
huffman 
parents: 
44282 
diff
changeset
 | 
231  | 
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
 | 
232  | 
using assms[of S] by simp  | 
| 40859 | 233  | 
qed  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
234  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
235  | 
lemma borel_measurable_const:  | 
| 38656 | 236  | 
"(\<lambda>x. c) \<in> borel_measurable M"  | 
| 47694 | 237  | 
by auto  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
238  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
239  | 
lemma borel_measurable_indicator:  | 
| 38656 | 240  | 
assumes A: "A \<in> sets M"  | 
241  | 
shows "indicator A \<in> borel_measurable M"  | 
|
| 46905 | 242  | 
unfolding indicator_def [abs_def] using A  | 
| 47694 | 243  | 
by (auto intro!: measurable_If_set)  | 
| 
33533
 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 
paulson 
parents:  
diff
changeset
 | 
244  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
245  | 
lemma borel_measurable_count_space[measurable (raw)]:  | 
| 50096 | 246  | 
"f \<in> borel_measurable (count_space S)"  | 
247  | 
unfolding measurable_def by auto  | 
|
248  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
249  | 
lemma borel_measurable_indicator'[measurable (raw)]:  | 
| 50096 | 250  | 
  assumes [measurable]: "{x\<in>space M. f x \<in> A x} \<in> sets M"
 | 
251  | 
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
 | 
252  | 
unfolding indicator_def[abs_def]  | 
| 
 
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
 
hoelzl 
parents: 
49774 
diff
changeset
 | 
253  | 
by (auto intro!: measurable_If)  | 
| 
 
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
 
hoelzl 
parents: 
49774 
diff
changeset
 | 
254  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
255  | 
lemma borel_measurable_indicator_iff:  | 
| 40859 | 256  | 
  "(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M"
 | 
257  | 
(is "?I \<in> borel_measurable M \<longleftrightarrow> _")  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
258  | 
proof  | 
| 40859 | 259  | 
assume "?I \<in> borel_measurable M"  | 
260  | 
  then have "?I -` {1} \<inter> space M \<in> sets M"
 | 
|
261  | 
unfolding measurable_def by auto  | 
|
262  | 
  also have "?I -` {1} \<inter> space M = A \<inter> space M"
 | 
|
| 46905 | 263  | 
unfolding indicator_def [abs_def] by auto  | 
| 40859 | 264  | 
finally show "A \<inter> space M \<in> sets M" .  | 
265  | 
next  | 
|
266  | 
assume "A \<inter> space M \<in> sets M"  | 
|
267  | 
moreover have "?I \<in> borel_measurable M \<longleftrightarrow>  | 
|
268  | 
(indicator (A \<inter> space M) :: 'a \<Rightarrow> 'x) \<in> borel_measurable M"  | 
|
269  | 
by (intro measurable_cong) (auto simp: indicator_def)  | 
|
270  | 
ultimately show "?I \<in> borel_measurable M" by auto  | 
|
271  | 
qed  | 
|
272  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
273  | 
lemma borel_measurable_subalgebra:  | 
| 41545 | 274  | 
assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N"  | 
| 39092 | 275  | 
shows "f \<in> borel_measurable M"  | 
276  | 
using assms unfolding measurable_def by auto  | 
|
277  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
278  | 
lemma borel_measurable_restrict_space_iff_ereal:  | 
| 57137 | 279  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
280  | 
assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"  | 
|
281  | 
shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>  | 
|
282  | 
(\<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
 | 
283  | 
by (subst measurable_restrict_space_iff)  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
284  | 
(auto simp: indicator_def of_bool_def if_distrib[where f="\<lambda>x. a * x" for a] cong: if_cong)  | 
| 57137 | 285  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
286  | 
lemma borel_measurable_restrict_space_iff_ennreal:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
287  | 
fixes f :: "'a \<Rightarrow> ennreal"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
288  | 
assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
289  | 
shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
290  | 
(\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
291  | 
by (subst measurable_restrict_space_iff)  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
292  | 
(auto simp: indicator_def of_bool_def if_distrib[where f="\<lambda>x. a * x" for a] cong: if_cong)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
293  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
294  | 
lemma borel_measurable_restrict_space_iff:  | 
| 57137 | 295  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
296  | 
assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"  | 
|
297  | 
shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>  | 
|
298  | 
(\<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
 | 
299  | 
by (subst measurable_restrict_space_iff)  | 
| 73536 | 300  | 
(auto simp: indicator_def of_bool_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] ac_simps  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
301  | 
cong: if_cong)  | 
| 
57138
 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
302  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
303  | 
lemma cbox_borel[measurable]: "cbox a b \<in> sets borel"  | 
| 
57138
 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
304  | 
by (auto intro: borel_closed)  | 
| 
 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
305  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
306  | 
lemma box_borel[measurable]: "box a b \<in> sets borel"  | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57275 
diff
changeset
 | 
307  | 
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
 | 
308  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
309  | 
lemma borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
310  | 
by (simp add: borel_closed compact_imp_closed)  | 
| 57137 | 311  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
312  | 
lemma borel_sigma_sets_subset:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
313  | 
"A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
314  | 
using sets.sigma_sets_subset[of A borel] by simp  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
315  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
316  | 
lemma borel_eq_sigmaI1:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
317  | 
fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
318  | 
assumes borel_eq: "borel = sigma UNIV X"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
319  | 
assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (F ` A))"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
320  | 
assumes F: "\<And>i. i \<in> A \<Longrightarrow> F i \<in> sets borel"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
321  | 
shows "borel = sigma UNIV (F ` A)"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
322  | 
unfolding borel_def  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
323  | 
proof (intro sigma_eqI antisym)  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
324  | 
  have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
325  | 
unfolding borel_def by simp  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
326  | 
also have "\<dots> = sigma_sets UNIV X"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
327  | 
unfolding borel_eq by simp  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
328  | 
also have "\<dots> \<subseteq> sigma_sets UNIV (F`A)"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
329  | 
using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
330  | 
  finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (F`A)" .
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
331  | 
  show "sigma_sets UNIV (F`A) \<subseteq> sigma_sets UNIV {S. open S}"
 | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
332  | 
by (metis F image_subset_iff sets_borel sigma_sets_mono)  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
333  | 
qed auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
334  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
335  | 
lemma borel_eq_sigmaI2:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
336  | 
fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
337  | 
and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
338  | 
assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`B)"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
339  | 
assumes X: "\<And>i j. (i, j) \<in> B \<Longrightarrow> G i j \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
340  | 
assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
341  | 
shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
342  | 
using assms  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
343  | 
by (smt (verit, del_insts) borel_eq_sigmaI1 image_iff prod.collapse split_beta)  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
344  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
345  | 
lemma borel_eq_sigmaI3:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
346  | 
fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
347  | 
assumes borel_eq: "borel = sigma UNIV X"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
348  | 
assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
349  | 
assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
350  | 
shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
351  | 
using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
352  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
353  | 
lemma borel_eq_sigmaI4:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
354  | 
fixes F :: "'i \<Rightarrow> 'a::topological_space set"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
355  | 
and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
356  | 
assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`A)"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
357  | 
assumes X: "\<And>i j. (i, j) \<in> A \<Longrightarrow> G i j \<in> sets (sigma UNIV (range F))"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
358  | 
assumes F: "\<And>i. F i \<in> sets borel"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
359  | 
shows "borel = sigma UNIV (range F)"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
360  | 
using assms by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` A" and F=F]) auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
361  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
362  | 
lemma borel_eq_sigmaI5:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
363  | 
fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
364  | 
assumes borel_eq: "borel = sigma UNIV (range G)"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
365  | 
assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
366  | 
assumes F: "\<And>i j. F i j \<in> sets borel"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
367  | 
shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
368  | 
using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
369  | 
|
| 
69722
 
b5163b2132c5
minor tagging updates in 13 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
370  | 
theorem second_countable_borel_measurable:  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
371  | 
fixes X :: "'a::second_countable_topology set set"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
372  | 
assumes eq: "open = generate_topology X"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
373  | 
shows "borel = sigma UNIV X"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
374  | 
unfolding borel_def  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
375  | 
proof (intro sigma_eqI sigma_sets_eqI)  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
376  | 
interpret X: sigma_algebra UNIV "sigma_sets UNIV X"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
377  | 
by (rule sigma_algebra_sigma_sets) simp  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
378  | 
|
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
379  | 
fix S :: "'a set" assume "S \<in> Collect open"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
380  | 
then have "generate_topology X S"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
381  | 
by (auto simp: eq)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
382  | 
then show "S \<in> sigma_sets UNIV X"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
383  | 
proof induction  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
384  | 
case (UN K)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
385  | 
then have K: "\<And>k. k \<in> K \<Longrightarrow> open k"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
386  | 
unfolding eq by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
387  | 
from ex_countable_basis obtain B :: "'a set set" where  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
388  | 
B: "\<And>b. b \<in> B \<Longrightarrow> open b" "\<And>X. open X \<Longrightarrow> \<exists>b\<subseteq>B. (\<Union>b) = X" and "countable B"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
389  | 
by (auto simp: topological_basis_def)  | 
| 69745 | 390  | 
from B(2)[OF K] obtain m where m: "\<And>k. k \<in> K \<Longrightarrow> m k \<subseteq> B" "\<And>k. k \<in> K \<Longrightarrow> \<Union>(m k) = k"  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
391  | 
by metis  | 
| 63040 | 392  | 
define U where "U = (\<Union>k\<in>K. m k)"  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
393  | 
with m have "countable U"  | 
| 61808 | 394  | 
by (intro countable_subset[OF _ \<open>countable B\<close>]) auto  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
395  | 
have "\<Union>U = (\<Union>A\<in>U. A)" by simp  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
396  | 
also have "\<dots> = \<Union>K"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
397  | 
unfolding U_def UN_simps by (simp add: m)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
398  | 
finally have "\<Union>U = \<Union>K" .  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
399  | 
|
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
400  | 
have "\<forall>b\<in>U. \<exists>k\<in>K. b \<subseteq> k"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
401  | 
using m by (auto simp: U_def)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
402  | 
then obtain u where u: "\<And>b. b \<in> U \<Longrightarrow> u b \<in> K" and "\<And>b. b \<in> U \<Longrightarrow> b \<subseteq> u b"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
403  | 
by metis  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
404  | 
then have "(\<Union>b\<in>U. u b) \<subseteq> \<Union>K" "\<Union>U \<subseteq> (\<Union>b\<in>U. u b)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
405  | 
by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
406  | 
then have "\<Union>K = (\<Union>b\<in>U. u b)"  | 
| 61808 | 407  | 
unfolding \<open>\<Union>U = \<Union>K\<close> by auto  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
408  | 
also have "\<dots> \<in> sigma_sets UNIV X"  | 
| 61808 | 409  | 
using u UN by (intro X.countable_UN' \<open>countable U\<close>) auto  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
410  | 
finally show "\<Union>K \<in> sigma_sets UNIV X" .  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
411  | 
qed auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
412  | 
qed (auto simp: eq intro: generate_topology.Basis)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
413  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
414  | 
lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
415  | 
proof -  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
416  | 
have "x \<in> sigma_sets UNIV (Collect closed)"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
417  | 
if "open x" for x :: "'a set"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
418  | 
by (metis that Compl_eq_Diff_UNIV closed_Compl double_complement mem_Collect_eq  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
419  | 
sigma_sets.Basic sigma_sets.Compl)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
420  | 
then show ?thesis  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
421  | 
unfolding borel_def  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
422  | 
by (metis Pow_UNIV borel_closed mem_Collect_eq sets_borel sigma_eqI sigma_sets_eqI top_greatest)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
423  | 
qed  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
424  | 
|
| 
69722
 
b5163b2132c5
minor tagging updates in 13 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
425  | 
proposition borel_eq_countable_basis:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
426  | 
fixes B::"'a::topological_space set set"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
427  | 
assumes "countable B"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
428  | 
assumes "topological_basis B"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
429  | 
shows "borel = sigma UNIV B"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
430  | 
unfolding borel_def  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
431  | 
proof (intro sigma_eqI sigma_sets_eqI, safe)  | 
| 69748 | 432  | 
interpret countable_basis "open" B using assms by (rule countable_basis_openI)  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
433  | 
fix X::"'a set" assume "open X"  | 
| 69748 | 434  | 
from open_countable_basisE[OF this] obtain B' where B': "B' \<subseteq> B" "X = \<Union> B'" .  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
435  | 
then show "X \<in> sigma_sets UNIV B"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
436  | 
by (blast intro: sigma_sets_UNION \<open>countable B\<close> countable_subset)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
437  | 
next  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
438  | 
fix b assume "b \<in> B"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
439  | 
hence "open b" by (rule topological_basis_open[OF assms(2)])  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
440  | 
thus "b \<in> sigma_sets UNIV (Collect open)" by auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
441  | 
qed simp_all  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
442  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
443  | 
lemma borel_measurable_continuous_on_restrict:  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
444  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"  | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
445  | 
assumes f: "continuous_on A f"  | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
446  | 
shows "f \<in> borel_measurable (restrict_space borel A)"  | 
| 
57138
 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
447  | 
proof (rule borel_measurableI)  | 
| 
 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
448  | 
fix S :: "'b set" assume "open S"  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
449  | 
with f obtain T where "f -` S \<inter> A = T \<inter> A" "open T"  | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
450  | 
by (metis continuous_on_open_invariant)  | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
451  | 
then show "f -` S \<inter> space (restrict_space borel A) \<in> sets (restrict_space borel A)"  | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
452  | 
by (force simp add: sets_restrict_space space_restrict_space)  | 
| 57137 | 453  | 
qed  | 
454  | 
||
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
455  | 
lemma borel_measurable_continuous_onI: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel"  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
456  | 
by (drule borel_measurable_continuous_on_restrict) simp  | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
457  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
458  | 
lemma borel_measurable_continuous_on_if:  | 
| 59415 | 459  | 
"A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> continuous_on (- A) g \<Longrightarrow>  | 
460  | 
(\<lambda>x. if x \<in> A then f x else g x) \<in> borel_measurable borel"  | 
|
461  | 
by (auto simp add: measurable_If_restrict_space_iff Collect_neg_eq  | 
|
462  | 
intro!: borel_measurable_continuous_on_restrict)  | 
|
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
463  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
464  | 
lemma borel_measurable_continuous_countable_exceptions:  | 
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57259 
diff
changeset
 | 
465  | 
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
 | 
466  | 
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
 | 
467  | 
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
 | 
468  | 
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
 | 
469  | 
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
 | 
470  | 
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
 | 
471  | 
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
 | 
472  | 
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
 | 
473  | 
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
 | 
474  | 
qed auto  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57259 
diff
changeset
 | 
475  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
476  | 
lemma borel_measurable_continuous_on:  | 
| 
57138
 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
477  | 
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
 | 
478  | 
shows "(\<lambda>x. f (g x)) \<in> borel_measurable M"  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
479  | 
using measurable_comp[OF g borel_measurable_continuous_onI[OF f]] by (simp add: comp_def)  | 
| 
57138
 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
480  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
481  | 
lemma borel_measurable_continuous_on_indicator:  | 
| 
57138
 
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
 
hoelzl 
parents: 
57137 
diff
changeset
 | 
482  | 
fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"  | 
| 59415 | 483  | 
shows "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable borel"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
484  | 
using borel_measurable_continuous_on_restrict borel_measurable_restrict_space_iff inf_top.right_neutral by blast  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
485  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
486  | 
lemma borel_measurable_Pair[measurable (raw)]:  | 
| 
50881
 
ae630bab13da
renamed countable_basis_space to second_countable_topology
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
487  | 
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
 | 
488  | 
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
 | 
489  | 
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
 | 
490  | 
shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
491  | 
proof (subst borel_eq_countable_basis)  | 
| 
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
 | 
492  | 
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
 | 
493  | 
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
 | 
494  | 
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
 | 
495  | 
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
 | 
496  | 
by (auto intro!: countable_basis topological_basis_prod is_basis)  | 
| 38656 | 497  | 
|
| 
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
 | 
498  | 
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
 | 
499  | 
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
 | 
500  | 
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
 | 
501  | 
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
 | 
502  | 
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
 | 
503  | 
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
 | 
504  | 
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
 | 
505  | 
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
 | 
506  | 
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
 | 
507  | 
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
 | 
508  | 
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
 | 
509  | 
qed auto  | 
| 
39087
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
510  | 
qed  | 
| 
 
96984bf6fa5b
Measurable on euclidean space is equiv. to measurable components
 
hoelzl 
parents: 
39083 
diff
changeset
 | 
511  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
512  | 
lemma borel_measurable_continuous_Pair:  | 
| 
50881
 
ae630bab13da
renamed countable_basis_space to second_countable_topology
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
513  | 
fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology"  | 
| 50003 | 514  | 
assumes [measurable]: "f \<in> borel_measurable M"  | 
515  | 
assumes [measurable]: "g \<in> borel_measurable M"  | 
|
| 49774 | 516  | 
assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))"  | 
517  | 
shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
518  | 
proof -  | 
| 49774 | 519  | 
have eq: "(\<lambda>x. H (f x) (g x)) = (\<lambda>x. (\<lambda>x. H (fst x) (snd x)) (f x, g x))" by auto  | 
520  | 
show ?thesis  | 
|
521  | 
unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto  | 
|
522  | 
qed  | 
|
523  | 
||
| 69683 | 524  | 
subsection \<open>Borel spaces on order topologies\<close>  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
525  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
526  | 
lemma [measurable]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
527  | 
fixes a b :: "'a::linorder_topology"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
528  | 
  shows lessThan_borel: "{..< a} \<in> sets borel"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
529  | 
    and greaterThan_borel: "{a <..} \<in> sets borel"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
530  | 
    and greaterThanLessThan_borel: "{a<..<b} \<in> sets borel"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
531  | 
    and atMost_borel: "{..a} \<in> sets borel"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
532  | 
    and atLeast_borel: "{a..} \<in> sets borel"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
533  | 
    and atLeastAtMost_borel: "{a..b} \<in> sets borel"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
534  | 
    and greaterThanAtMost_borel: "{a<..b} \<in> sets borel"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
535  | 
    and atLeastLessThan_borel: "{a..<b} \<in> sets borel"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
536  | 
unfolding greaterThanAtMost_def atLeastLessThan_def  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
537  | 
by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
538  | 
closed_atMost closed_atLeast closed_atLeastAtMost)+  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
539  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
540  | 
lemma borel_Iio:  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
541  | 
  "borel = sigma UNIV (range lessThan :: 'a::{linorder_topology, second_countable_topology} set set)"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
542  | 
unfolding second_countable_borel_measurable[OF open_generated_order]  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
543  | 
proof (intro sigma_eqI sigma_sets_eqI)  | 
| 74362 | 544  | 
  obtain D :: "'a set" where D: "countable D" "\<And>X. open X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> \<exists>d\<in>D. d \<in> X"
 | 
545  | 
by (rule countable_dense_setE) blast  | 
|
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
546  | 
|
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
547  | 
interpret L: sigma_algebra UNIV "sigma_sets UNIV (range lessThan)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
548  | 
by (rule sigma_algebra_sigma_sets) simp  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
549  | 
|
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
550  | 
fix A :: "'a set" assume "A \<in> range lessThan \<union> range greaterThan"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
551  | 
  then obtain y where "A = {y <..} \<or> A = {..< y}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
552  | 
by blast  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
553  | 
then show "A \<in> sigma_sets UNIV (range lessThan)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
554  | 
proof  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
555  | 
    assume A: "A = {y <..}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
556  | 
show ?thesis  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
557  | 
proof cases  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
558  | 
assume "\<forall>x>y. \<exists>d. y < d \<and> d < x"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
559  | 
      with D(2)[of "{y <..< x}" for x] have "\<forall>x>y. \<exists>d\<in>D. y < d \<and> d < x"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
560  | 
by (auto simp: set_eq_iff)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
561  | 
      then have "A = UNIV - (\<Inter>d\<in>{d\<in>D. y < d}. {..< d})"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
562  | 
by (auto simp: A) (metis less_asym)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
563  | 
also have "\<dots> \<in> sigma_sets UNIV (range lessThan)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
564  | 
using D(1) by (intro L.Diff L.top L.countable_INT'') auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
565  | 
finally show ?thesis .  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
566  | 
next  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
567  | 
assume "\<not> (\<forall>x>y. \<exists>d. y < d \<and> d < x)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
568  | 
then obtain x where "y < x" "\<And>d. y < d \<Longrightarrow> \<not> d < x"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
569  | 
by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
570  | 
      then have "A = UNIV - {..< x}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
571  | 
unfolding A by (auto simp: not_less[symmetric])  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
572  | 
also have "\<dots> \<in> sigma_sets UNIV (range lessThan)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
573  | 
by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
574  | 
finally show ?thesis .  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
575  | 
qed  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
576  | 
qed auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
577  | 
qed auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
578  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
579  | 
lemma borel_Ioi:  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
580  | 
  "borel = sigma UNIV (range greaterThan :: 'a::{linorder_topology, second_countable_topology} set set)"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
581  | 
unfolding second_countable_borel_measurable[OF open_generated_order]  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
582  | 
proof (intro sigma_eqI sigma_sets_eqI)  | 
| 74362 | 583  | 
  obtain D :: "'a set" where D: "countable D" "\<And>X. open X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> \<exists>d\<in>D. d \<in> X"
 | 
584  | 
by (rule countable_dense_setE) blast  | 
|
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
585  | 
|
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
586  | 
interpret L: sigma_algebra UNIV "sigma_sets UNIV (range greaterThan)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
587  | 
by (rule sigma_algebra_sigma_sets) simp  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
588  | 
|
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
589  | 
fix A :: "'a set" assume "A \<in> range lessThan \<union> range greaterThan"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
590  | 
  then obtain y where "A = {y <..} \<or> A = {..< y}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
591  | 
by blast  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
592  | 
then show "A \<in> sigma_sets UNIV (range greaterThan)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
593  | 
proof  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
594  | 
    assume A: "A = {..< y}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
595  | 
show ?thesis  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
596  | 
proof cases  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
597  | 
assume "\<forall>x<y. \<exists>d. x < d \<and> d < y"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
598  | 
      with D(2)[of "{x <..< y}" for x] have "\<forall>x<y. \<exists>d\<in>D. x < d \<and> d < y"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
599  | 
by (auto simp: set_eq_iff)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
600  | 
      then have "A = UNIV - (\<Inter>d\<in>{d\<in>D. d < y}. {d <..})"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
601  | 
by (auto simp: A) (metis less_asym)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
602  | 
also have "\<dots> \<in> sigma_sets UNIV (range greaterThan)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
603  | 
using D(1) by (intro L.Diff L.top L.countable_INT'') auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
604  | 
finally show ?thesis .  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
605  | 
next  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
606  | 
assume "\<not> (\<forall>x<y. \<exists>d. x < d \<and> d < y)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
607  | 
then obtain x where "x < y" "\<And>d. y > d \<Longrightarrow> x \<ge> d"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
608  | 
by (auto simp: not_less[symmetric])  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
609  | 
      then have "A = UNIV - {x <..}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
610  | 
unfolding A Compl_eq_Diff_UNIV[symmetric] by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
611  | 
also have "\<dots> \<in> sigma_sets UNIV (range greaterThan)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
612  | 
by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
613  | 
finally show ?thesis .  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
614  | 
qed  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
615  | 
qed auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
616  | 
qed auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
617  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
618  | 
lemma borel_measurableI_less:  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
619  | 
  fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
620  | 
  shows "(\<And>y. {x\<in>space M. f x < y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
621  | 
unfolding borel_Iio  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
622  | 
by (rule measurable_measure_of) (auto simp: Int_def conj_commute)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
623  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
624  | 
lemma borel_measurableI_greater:  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
625  | 
  fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
626  | 
  shows "(\<And>y. {x\<in>space M. y < f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
627  | 
unfolding borel_Ioi  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
628  | 
by (rule measurable_measure_of) (auto simp: Int_def conj_commute)  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
629  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
630  | 
lemma borel_measurableI_le:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
631  | 
  fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
632  | 
  shows "(\<And>y. {x\<in>space M. f x \<le> y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
633  | 
by (rule borel_measurableI_greater) (auto simp: not_le[symmetric])  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
634  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
635  | 
lemma borel_measurableI_ge:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
636  | 
  fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
637  | 
  shows "(\<And>y. {x\<in>space M. y \<le> f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
638  | 
by (rule borel_measurableI_less) (auto simp: not_le[symmetric])  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
639  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
640  | 
lemma borel_measurable_less[measurable]:  | 
| 63332 | 641  | 
  fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
642  | 
assumes "f \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
643  | 
assumes "g \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
644  | 
  shows "{w \<in> space M. f w < g w} \<in> sets M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
645  | 
proof -  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
646  | 
  have "{w \<in> space M. f w < g w} = (\<lambda>x. (f x, g x)) -` {x. fst x < snd x} \<inter> space M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
647  | 
by auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
648  | 
also have "\<dots> \<in> sets M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
649  | 
by (intro measurable_sets[OF borel_measurable_Pair borel_open, OF assms open_Collect_less]  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
650  | 
continuous_intros)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
651  | 
finally show ?thesis .  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
652  | 
qed  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
653  | 
|
| 69739 | 654  | 
lemma  | 
| 63332 | 655  | 
  fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
656  | 
assumes f[measurable]: "f \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
657  | 
assumes g[measurable]: "g \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
658  | 
  shows borel_measurable_le[measurable]: "{w \<in> space M. f w \<le> g w} \<in> sets M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
659  | 
    and borel_measurable_eq[measurable]: "{w \<in> space M. f w = g w} \<in> sets M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
660  | 
    and borel_measurable_neq: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
661  | 
unfolding eq_iff not_less[symmetric]  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
662  | 
by measurable  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
663  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
664  | 
lemma borel_measurable_SUP[measurable (raw)]:  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
665  | 
  fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
666  | 
assumes [simp]: "countable I"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
667  | 
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
668  | 
shows "(\<lambda>x. SUP i\<in>I. F i x) \<in> borel_measurable M"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
669  | 
by (rule borel_measurableI_greater) (simp add: less_SUP_iff)  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
670  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
671  | 
lemma borel_measurable_INF[measurable (raw)]:  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
672  | 
  fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
673  | 
assumes [simp]: "countable I"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
674  | 
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
675  | 
shows "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable M"  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
676  | 
by (rule borel_measurableI_less) (simp add: INF_less_iff)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
677  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
678  | 
lemma borel_measurable_cSUP[measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
679  | 
  fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
680  | 
assumes [simp]: "countable I"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
681  | 
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
682  | 
assumes bdd: "\<And>x. x \<in> space M \<Longrightarrow> bdd_above ((\<lambda>i. F i x) ` I)"  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
683  | 
shows "(\<lambda>x. SUP i\<in>I. F i x) \<in> borel_measurable M"  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
684  | 
proof cases  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
685  | 
  assume "I = {}" then show ?thesis
 | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
686  | 
by (simp add: borel_measurable_const)  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
687  | 
next  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
688  | 
  assume "I \<noteq> {}"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
689  | 
show ?thesis  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
690  | 
proof (rule borel_measurableI_le)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
691  | 
fix y  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
692  | 
    have "{x \<in> space M. \<forall>i\<in>I. F i x \<le> y} \<in> sets M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
693  | 
by measurable  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
694  | 
    also have "{x \<in> space M. \<forall>i\<in>I. F i x \<le> y} = {x \<in> space M. (SUP i\<in>I. F i x) \<le> y}"
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
695  | 
      by (simp add: cSUP_le_iff \<open>I \<noteq> {}\<close> bdd cong: conj_cong)
 | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
696  | 
    finally show "{x \<in> space M. (SUP i\<in>I. F i x) \<le>  y} \<in> sets M"  .
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
697  | 
qed  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
698  | 
qed  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
699  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
700  | 
lemma borel_measurable_cINF[measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
701  | 
  fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
702  | 
assumes [simp]: "countable I"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
703  | 
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
704  | 
assumes bdd: "\<And>x. x \<in> space M \<Longrightarrow> bdd_below ((\<lambda>i. F i x) ` I)"  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
705  | 
shows "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable M"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
706  | 
proof cases  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
707  | 
  assume "I = {}" then show ?thesis
 | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
708  | 
by (simp add: borel_measurable_const)  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
709  | 
next  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
710  | 
  assume "I \<noteq> {}"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
711  | 
show ?thesis  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
712  | 
proof (rule borel_measurableI_ge)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
713  | 
fix y  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
714  | 
    have "{x \<in> space M. \<forall>i\<in>I. y \<le> F i x} \<in> sets M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
715  | 
by measurable  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
716  | 
    also have "{x \<in> space M. \<forall>i\<in>I. y \<le> F i x} = {x \<in> space M. y \<le> (INF i\<in>I. F i x)}"
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
717  | 
      by (simp add: le_cINF_iff \<open>I \<noteq> {}\<close> bdd cong: conj_cong)
 | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
718  | 
    finally show "{x \<in> space M. y \<le> (INF i\<in>I. F i x)} \<in> sets M"  .
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
719  | 
qed  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
720  | 
qed  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
721  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
722  | 
lemma borel_measurable_lfp[consumes 1, case_names continuity step]:  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
723  | 
  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})"
 | 
| 
60172
 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 
hoelzl 
parents: 
60150 
diff
changeset
 | 
724  | 
assumes "sup_continuous F"  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
725  | 
assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
726  | 
shows "lfp F \<in> borel_measurable M"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
727  | 
proof -  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
728  | 
  { fix i have "((F ^^ i) bot) \<in> borel_measurable M"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
729  | 
by (induct i) (auto intro!: *) }  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
730  | 
then have "(\<lambda>x. SUP i. (F ^^ i) bot x) \<in> borel_measurable M"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
731  | 
by measurable  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
732  | 
also have "(\<lambda>x. SUP i. (F ^^ i) bot x) = (SUP i. (F ^^ i) bot)"  | 
| 
69861
 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 
haftmann 
parents: 
69748 
diff
changeset
 | 
733  | 
by (auto simp add: image_comp)  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
734  | 
also have "(SUP i. (F ^^ i) bot) = lfp F"  | 
| 
60172
 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 
hoelzl 
parents: 
60150 
diff
changeset
 | 
735  | 
by (rule sup_continuous_lfp[symmetric]) fact  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
736  | 
finally show ?thesis .  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
737  | 
qed  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
738  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
739  | 
lemma borel_measurable_gfp[consumes 1, case_names continuity step]:  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
740  | 
  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})"
 | 
| 
60172
 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 
hoelzl 
parents: 
60150 
diff
changeset
 | 
741  | 
assumes "inf_continuous F"  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
742  | 
assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
743  | 
shows "gfp F \<in> borel_measurable M"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
744  | 
proof -  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
745  | 
  { fix i have "((F ^^ i) top) \<in> borel_measurable M"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
746  | 
by (induct i) (auto intro!: * simp: bot_fun_def) }  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
747  | 
then have "(\<lambda>x. INF i. (F ^^ i) top x) \<in> borel_measurable M"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
748  | 
by measurable  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
749  | 
also have "(\<lambda>x. INF i. (F ^^ i) top x) = (INF i. (F ^^ i) top)"  | 
| 
69861
 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 
haftmann 
parents: 
69748 
diff
changeset
 | 
750  | 
by (auto simp add: image_comp)  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
751  | 
also have "\<dots> = gfp F"  | 
| 
60172
 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 
hoelzl 
parents: 
60150 
diff
changeset
 | 
752  | 
by (rule inf_continuous_gfp[symmetric]) fact  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
753  | 
finally show ?thesis .  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
754  | 
qed  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59000 
diff
changeset
 | 
755  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
756  | 
lemma borel_measurable_max[measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
757  | 
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. max (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
758  | 
by (rule borel_measurableI_less) simp  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
759  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
760  | 
lemma borel_measurable_min[measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
761  | 
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. min (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
762  | 
by (rule borel_measurableI_greater) simp  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
763  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
764  | 
lemma borel_measurable_Min[measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
765  | 
  "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, linorder_topology}) \<in> borel_measurable M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
766  | 
proof (induct I rule: finite_induct)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
767  | 
case (insert i I) then show ?case  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
768  | 
    by (cases "I = {}") auto
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
769  | 
qed auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
770  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
771  | 
lemma borel_measurable_Max[measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
772  | 
  "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, linorder_topology}) \<in> borel_measurable M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
773  | 
proof (induct I rule: finite_induct)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
774  | 
case (insert i I) then show ?case  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
775  | 
    by (cases "I = {}") auto
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
776  | 
qed auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
777  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
778  | 
lemma borel_measurable_sup[measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
779  | 
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. sup (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
780  | 
unfolding sup_max by measurable  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
781  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
782  | 
lemma borel_measurable_inf[measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
783  | 
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. inf (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
784  | 
unfolding inf_min by measurable  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
785  | 
|
| 69739 | 786  | 
lemma [measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
787  | 
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
788  | 
assumes "\<And>i. f i \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
789  | 
shows borel_measurable_liminf: "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
790  | 
and borel_measurable_limsup: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
791  | 
unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
792  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
793  | 
lemma measurable_convergent[measurable (raw)]:  | 
| 63332 | 794  | 
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
795  | 
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
796  | 
shows "Measurable.pred M (\<lambda>x. convergent (\<lambda>i. f i x))"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
797  | 
unfolding convergent_ereal by measurable  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
798  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
799  | 
lemma sets_Collect_convergent[measurable]:  | 
| 63332 | 800  | 
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
801  | 
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
802  | 
  shows "{x\<in>space M. convergent (\<lambda>i. f i x)} \<in> sets M"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
803  | 
by measurable  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
804  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
805  | 
lemma borel_measurable_lim[measurable (raw)]:  | 
| 63332 | 806  | 
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
807  | 
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
808  | 
shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
809  | 
proof -  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
810  | 
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))"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
811  | 
by (simp add: lim_def convergent_def convergent_limsup_cl)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
812  | 
then show ?thesis  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
813  | 
by simp  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
814  | 
qed  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
815  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
816  | 
lemma borel_measurable_LIMSEQ_order:  | 
| 63332 | 817  | 
  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
818  | 
assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
819  | 
and u: "\<And>i. u i \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
820  | 
shows "u' \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
821  | 
proof -  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
822  | 
have "\<And>x. x \<in> space M \<Longrightarrow> u' x = liminf (\<lambda>n. u n x)"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
823  | 
using u' by (simp add: lim_imp_Liminf[symmetric])  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
824  | 
with u show ?thesis by (simp cong: measurable_cong)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
825  | 
qed  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
826  | 
|
| 69683 | 827  | 
subsection \<open>Borel spaces on topological monoids\<close>  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
828  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
829  | 
lemma borel_measurable_add[measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
830  | 
  fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, topological_monoid_add}"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
831  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
832  | 
assumes g: "g \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
833  | 
shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
834  | 
using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
835  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
836  | 
lemma borel_measurable_sum[measurable (raw)]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
837  | 
  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, topological_comm_monoid_add}"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
838  | 
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
839  | 
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
840  | 
proof cases  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
841  | 
assume "finite S"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
842  | 
thus ?thesis using assms by induct auto  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
843  | 
qed simp  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
844  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
845  | 
lemma borel_measurable_suminf_order[measurable (raw)]:  | 
| 63332 | 846  | 
  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology, topological_comm_monoid_add}"
 | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
847  | 
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
848  | 
shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
849  | 
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
850  | 
|
| 69683 | 851  | 
subsection \<open>Borel spaces on Euclidean spaces\<close>  | 
| 
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  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
853  | 
lemma borel_measurable_inner[measurable (raw)]:  | 
| 
50881
 
ae630bab13da
renamed countable_basis_space to second_countable_topology
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
854  | 
  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
 | 
855  | 
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
 | 
856  | 
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
 | 
857  | 
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
 | 
858  | 
using assms  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
859  | 
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
 | 
860  | 
|
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
861  | 
notation  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
862  | 
eucl_less (infix "<e" 50)  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
863  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
864  | 
lemma box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}"
 | 
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
865  | 
  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
 | 
866  | 
by auto  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
867  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
868  | 
lemma eucl_ivals[measurable]:  | 
| 61076 | 869  | 
fixes a b :: "'a::ordered_euclidean_space"  | 
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
870  | 
  shows "{x. x <e a} \<in> sets borel"
 | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
871  | 
    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
 | 
872  | 
    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
 | 
873  | 
    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
 | 
874  | 
    and "{a..b} \<in> sets borel"
 | 
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
875  | 
    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
 | 
876  | 
    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
 | 
877  | 
unfolding box_oc box_co  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
878  | 
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
 | 
879  | 
|
| 69739 | 880  | 
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
 | 
881  | 
  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
 | 
882  | 
  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
 | 
883  | 
    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
 | 
884  | 
    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
 | 
885  | 
    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
 | 
886  | 
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
 | 
887  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
888  | 
lemma borel_eq_box:  | 
| 61076 | 889  | 
"borel = sigma UNIV (range (\<lambda> (a, b). box a b :: 'a :: euclidean_space set))"  | 
| 
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
 | 
890  | 
(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
 | 
891  | 
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
 | 
892  | 
  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
 | 
893  | 
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
 | 
894  | 
show "M \<in> ?SIGMA"  | 
| 61808 | 895  | 
apply (subst open_UNION_box[OF \<open>open M\<close>])  | 
| 
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
 | 
896  | 
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
 | 
897  | 
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
 | 
898  | 
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
 | 
899  | 
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
 | 
900  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
901  | 
lemma halfspace_gt_in_halfspace:  | 
| 
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
 | 
902  | 
assumes i: "i \<in> A"  | 
| 62372 | 903  | 
  shows "{x::'a. a < x \<bullet> i} \<in>
 | 
| 61076 | 904  | 
    sigma_sets UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> 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
 | 
905  | 
(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
 | 
906  | 
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
 | 
907  | 
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
 | 
908  | 
by (intro sigma_algebra_sigma_sets) simp_all  | 
| 61076 | 909  | 
  have *: "?set = (\<Union>n. UNIV - {x::'a. x \<bullet> i < a + 1 / real (Suc n)})"
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
910  | 
proof (safe, simp_all add: not_less del: of_nat_Suc)  | 
| 
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
 | 
911  | 
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
 | 
912  | 
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
 | 
913  | 
obtain n where "a + 1 / real (Suc n) < x \<bullet> i"  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
914  | 
by (auto simp: field_simps)  | 
| 
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
 | 
915  | 
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
 | 
916  | 
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
 | 
917  | 
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
 | 
918  | 
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
 | 
919  | 
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
 | 
920  | 
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
 | 
921  | 
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
 | 
922  | 
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
 | 
923  | 
show "?set \<in> ?SIGMA" unfolding *  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61284 
diff
changeset
 | 
924  | 
by (auto intro!: Diff sigma_sets_Inter i)  | 
| 
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
 | 
925  | 
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
 | 
926  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
927  | 
lemma borel_eq_halfspace_less:  | 
| 
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
 | 
928  | 
  "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
 | 
929  | 
(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
 | 
930  | 
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
 | 
931  | 
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
 | 
932  | 
  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
 | 
933  | 
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
 | 
934  | 
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
 | 
935  | 
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
 | 
936  | 
(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
 | 
937  | 
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
 | 
938  | 
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
 | 
939  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
940  | 
lemma borel_eq_halfspace_le:  | 
| 
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
 | 
941  | 
  "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
 | 
942  | 
(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
 | 
943  | 
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
 | 
944  | 
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
 | 
945  | 
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
 | 
946  | 
  have *: "{x::'a. x\<bullet>i < a} = (\<Union>n. {x. x\<bullet>i \<le> a - 1/real (Suc n)})"
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
947  | 
proof (safe, simp_all del: of_nat_Suc)  | 
| 
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
 | 
948  | 
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
 | 
949  | 
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
 | 
950  | 
obtain n where "x \<bullet> i < a - 1 / (real (Suc n))"  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
951  | 
by (auto simp: field_simps)  | 
| 
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
 | 
952  | 
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
 | 
953  | 
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
 | 
954  | 
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
 | 
955  | 
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
 | 
956  | 
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
 | 
957  | 
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
 | 
958  | 
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
 | 
959  | 
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
 | 
960  | 
  show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding *
 | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
961  | 
by (intro sets.countable_UN) (auto intro: i)  | 
| 
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
 | 
962  | 
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
 | 
963  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
964  | 
lemma borel_eq_halfspace_ge:  | 
| 61076 | 965  | 
  "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a \<le> x \<bullet> i}) ` (UNIV \<times> Basis))"
 | 
| 
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
 | 
966  | 
(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
 | 
967  | 
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
 | 
968  | 
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
 | 
969  | 
  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
 | 
970  | 
  show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding *
 | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
971  | 
using i by (intro sets.compl_sets) auto  | 
| 
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
 | 
972  | 
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
 | 
973  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
974  | 
lemma borel_eq_halfspace_greater:  | 
| 61076 | 975  | 
  "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))"
 | 
| 
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
 | 
976  | 
(is "_ = ?SIGMA")  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
977  | 
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])  | 
| 
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
 | 
978  | 
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
 | 
979  | 
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
 | 
980  | 
  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
 | 
981  | 
  show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding *
 | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
982  | 
by (intro sets.compl_sets) (auto intro: i)  | 
| 
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
 | 
983  | 
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
 | 
984  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
985  | 
lemma borel_eq_atMost:  | 
| 61076 | 986  | 
  "borel = sigma UNIV (range (\<lambda>a. {..a::'a::ordered_euclidean_space}))"
 | 
| 
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
 | 
987  | 
(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
 | 
988  | 
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
 | 
989  | 
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
 | 
990  | 
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
 | 
991  | 
  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)})"
 | 
| 62390 | 992  | 
proof (safe, simp_all add: eucl_le[where 'a='a] split: if_split_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
 | 
993  | 
fix x :: 'a  | 
| 74362 | 994  | 
obtain k where "Max ((\<bullet>) x ` Basis) \<le> real k"  | 
995  | 
using real_arch_simple by blast  | 
|
| 
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
 | 
996  | 
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
 | 
997  | 
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
 | 
998  | 
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
 | 
999  | 
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
 | 
1000  | 
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
 | 
1001  | 
  show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding *
 | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1002  | 
by (intro sets.countable_UN) auto  | 
| 
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
 | 
1003  | 
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
 | 
1004  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1005  | 
lemma borel_eq_greaterThan:  | 
| 61076 | 1006  | 
  "borel = sigma UNIV (range (\<lambda>a::'a::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
 | 
1007  | 
(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
 | 
1008  | 
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
 | 
1009  | 
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
 | 
1010  | 
then have i: "i \<in> Basis" by auto  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1011  | 
have **: "\<exists>y. \<forall>j\<in>Basis. j \<noteq> i \<longrightarrow> - real y < x \<bullet> j" if "a < x \<bullet> i" for x  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1012  | 
proof -  | 
| 74362 | 1013  | 
obtain k where k: "Max ((\<bullet>) (- x) ` Basis) < real k"  | 
1014  | 
using reals_Archimedean2 by blast  | 
|
| 
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
 | 
1015  | 
    { 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
 | 
1016  | 
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
 | 
1017  | 
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
 | 
1018  | 
then have "- real k < x\<bullet>i" by simp }  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1019  | 
then 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
 | 
1020  | 
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
 | 
1021  | 
qed  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1022  | 
  have "{x::'a. x\<bullet>i \<le> a} = UNIV - {x::'a. a < x\<bullet>i}" by auto
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1023  | 
  also have *: "{x::'a. a < x\<bullet>i} = (\<Union>k::nat. {x. (\<Sum>n\<in>Basis. (if n = i then a else -k) *\<^sub>R n) <e x})" 
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1024  | 
using i ** by (force simp add: eucl_less_def split: if_split_asm)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1025  | 
  finally have eq: "{x. x \<bullet> i \<le> a} = UNIV - (\<Union>x. {xa. (\<Sum>n\<in>Basis. (if n = i then a else - real x) *\<^sub>R n) <e xa})" .
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1026  | 
  show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA"
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1027  | 
unfolding eq by (fastforce intro!: sigma_sets_top sets.Diff)  | 
| 
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
 | 
1028  | 
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
 | 
1029  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1030  | 
lemma borel_eq_lessThan:  | 
| 61076 | 1031  | 
  "borel = sigma UNIV (range (\<lambda>a::'a::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
 | 
1032  | 
(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
 | 
1033  | 
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
 | 
1034  | 
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
 | 
1035  | 
then have i: "i \<in> Basis" by auto  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1036  | 
have **: "\<exists>y. \<forall>j\<in>Basis. j \<noteq> i \<longrightarrow> real y > x \<bullet> j" if "a > x \<bullet> i" for x  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1037  | 
proof -  | 
| 74362 | 1038  | 
obtain k where k: "Max ((\<bullet>) x ` Basis) < real k"  | 
1039  | 
using reals_Archimedean2 by blast  | 
|
| 
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
 | 
1040  | 
    { 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
 | 
1041  | 
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
 | 
1042  | 
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
 | 
1043  | 
then have "x\<bullet>i < real k" by simp }  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1044  | 
then 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
 | 
1045  | 
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
 | 
1046  | 
qed  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1047  | 
  have "{x::'a. a \<le> x\<bullet>i} = UNIV - {x::'a. x\<bullet>i < a}" by auto
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1048  | 
  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 \<open>i\<in> Basis\<close>
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1049  | 
using i ** by (force simp add: eucl_less_def split: if_split_asm)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1050  | 
finally  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1051  | 
  have eq: "{x. a \<le> x \<bullet> i} =
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1052  | 
            UNIV - (\<Union>k. {x. x <e (\<Sum>n\<in>Basis. (if n = i then a else real k) *\<^sub>R n)})" .
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1053  | 
|
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1054  | 
  show "{x. a \<le> x\<bullet>i} \<in> ?SIGMA"
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1055  | 
unfolding eq by (fastforce intro!: sigma_sets_top sets.Diff)  | 
| 
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
 | 
1056  | 
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
 | 
1057  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1058  | 
lemma borel_eq_atLeastAtMost:  | 
| 61076 | 1059  | 
  "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} ::'a::ordered_euclidean_space set))"
 | 
| 
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
 | 
1060  | 
(is "_ = ?SIGMA")  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1061  | 
proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])  | 
| 
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
 | 
1062  | 
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
 | 
1063  | 
  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
 | 
1064  | 
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
 | 
1065  | 
fix x :: 'a  | 
| 74362 | 1066  | 
obtain k where k: "Max ((\<bullet>) (- x) ` Basis) \<le> real k"  | 
1067  | 
using real_arch_simple by blast  | 
|
| 
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
 | 
1068  | 
    { 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
 | 
1069  | 
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
 | 
1070  | 
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
 | 
1071  | 
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
 | 
1072  | 
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
 | 
1073  | 
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
 | 
1074  | 
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
 | 
1075  | 
  show "{..a} \<in> ?SIGMA" unfolding *
 | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1076  | 
by (intro sets.countable_UN)  | 
| 
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
 | 
1077  | 
(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
 | 
1078  | 
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
 | 
1079  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1080  | 
lemma borel_set_induct[consumes 1, case_names empty interval compl union]:  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1081  | 
assumes "A \<in> sets borel"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1082  | 
  assumes empty: "P {}" and int: "\<And>a b. a \<le> b \<Longrightarrow> P {a..b}" and compl: "\<And>A. A \<in> sets borel \<Longrightarrow> P A \<Longrightarrow> P (-A)" and
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1083  | 
un: "\<And>f. disjoint_family f \<Longrightarrow> (\<And>i. f i \<in> sets borel) \<Longrightarrow> (\<And>i. P (f i)) \<Longrightarrow> P (\<Union>i::nat. f i)"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1084  | 
shows "P (A::real set)"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1085  | 
proof -  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1086  | 
  let ?G = "range (\<lambda>(a,b). {a..b::real})"
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1087  | 
have "Int_stable ?G" "?G \<subseteq> Pow UNIV" "A \<in> sigma_sets UNIV ?G"  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1088  | 
using assms(1) by (auto simp add: borel_eq_atLeastAtMost Int_stable_def)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1089  | 
thus ?thesis  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1090  | 
proof (induction rule: sigma_sets_induct_disjoint)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1091  | 
case (union f)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1092  | 
from union.hyps(2) have "\<And>i. f i \<in> sets borel" by (auto simp: borel_eq_atLeastAtMost)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1093  | 
with union show ?case by (auto intro: un)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1094  | 
next  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1095  | 
case (basic A)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1096  | 
    then obtain a b where "A = {a .. b}" by auto
 | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1097  | 
then show ?case  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1098  | 
by (cases "a \<le> b") (auto intro: int empty)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1099  | 
qed (auto intro: empty compl simp: Compl_eq_Diff_UNIV[symmetric] borel_eq_atLeastAtMost)  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1100  | 
qed  | 
| 
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1101  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1102  | 
lemma borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))"
 | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57275 
diff
changeset
 | 
1103  | 
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
 | 
1104  | 
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
 | 
1105  | 
  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
 | 
1106  | 
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
 | 
1107  | 
  also have "\<dots> \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))"
 | 
| 62372 | 1108  | 
by (intro sets.countable_nat_UN) auto  | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57275 
diff
changeset
 | 
1109  | 
  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
 | 
1110  | 
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
 | 
1111  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1112  | 
lemma eucl_lessThan: "{x::real. x <e a} = lessThan a"
 | 
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
1113  | 
by (simp add: eucl_less_def lessThan_def)  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
1114  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1115  | 
lemma borel_eq_atLeastLessThan:  | 
| 
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
 | 
1116  | 
  "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
 | 
1117  | 
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
 | 
1118  | 
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
 | 
1119  | 
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
 | 
1120  | 
  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
 | 
1121  | 
by (auto simp: move_uminus real_arch_simple)  | 
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
1122  | 
  then show "{y. y <e x} \<in> ?SIGMA"
 | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1123  | 
by (auto intro: sigma_sets.intros(2-) 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
 | 
1124  | 
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
 | 
1125  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1126  | 
lemma borel_measurable_halfspacesI:  | 
| 61076 | 1127  | 
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"  | 
| 
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
 | 
1128  | 
assumes F: "borel = sigma UNIV (F ` (UNIV \<times> Basis))"  | 
| 62372 | 1129  | 
and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space 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
 | 
1130  | 
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
 | 
1131  | 
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
 | 
1132  | 
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
 | 
1133  | 
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
 | 
1134  | 
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
 | 
1135  | 
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
 | 
1136  | 
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
 | 
1137  | 
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
 | 
1138  | 
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
 | 
1139  | 
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
 | 
1140  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1141  | 
lemma borel_measurable_iff_halfspace_le:  | 
| 61076 | 1142  | 
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"  | 
| 
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
 | 
1143  | 
  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
 | 
1144  | 
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
 | 
1145  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1146  | 
lemma borel_measurable_iff_halfspace_less:  | 
| 61076 | 1147  | 
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"  | 
| 
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
 | 
1148  | 
  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
 | 
1149  | 
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
 | 
1150  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1151  | 
lemma borel_measurable_iff_halfspace_ge:  | 
| 61076 | 1152  | 
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"  | 
| 
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
 | 
1153  | 
  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
 | 
1154  | 
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
 | 
1155  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1156  | 
lemma borel_measurable_iff_halfspace_greater:  | 
| 61076 | 1157  | 
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"  | 
| 
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
 | 
1158  | 
  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
 | 
1159  | 
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
 | 
1160  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1161  | 
lemma borel_measurable_iff_le:  | 
| 
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
 | 
1162  | 
  "(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
 | 
1163  | 
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
 | 
1164  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1165  | 
lemma borel_measurable_iff_less:  | 
| 
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
 | 
1166  | 
  "(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
 | 
1167  | 
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
 | 
1168  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1169  | 
lemma borel_measurable_iff_ge:  | 
| 
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
 | 
1170  | 
  "(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
 | 
1171  | 
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
 | 
1172  | 
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
 | 
1173  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1174  | 
lemma borel_measurable_iff_greater:  | 
| 
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
 | 
1175  | 
  "(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
 | 
1176  | 
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
 | 
1177  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1178  | 
lemma borel_measurable_euclidean_space:  | 
| 
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
 | 
1179  | 
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
 | 
1180  | 
shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M)"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1181  | 
proof safe  | 
| 
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
 | 
1182  | 
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
 | 
1183  | 
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
 | 
1184  | 
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
 | 
1185  | 
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
 | 
1186  | 
|
| 69683 | 1187  | 
subsection "Borel measurable operators"  | 
| 
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
 | 
1188  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1189  | 
lemma borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel"  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1190  | 
by (intro borel_measurable_continuous_onI continuous_intros)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1191  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1192  | 
lemma borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel"  | 
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57259 
diff
changeset
 | 
1193  | 
  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
 | 
1194  | 
(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
 | 
1195  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1196  | 
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
 | 
1197  | 
  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
 | 
1198  | 
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
 | 
1199  | 
shows "(\<lambda>x. - g x) \<in> borel_measurable M"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1200  | 
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
 | 
1201  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1202  | 
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
 | 
1203  | 
  fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
 | 
| 49774 | 1204  | 
assumes f: "f \<in> borel_measurable M"  | 
1205  | 
assumes g: "g \<in> borel_measurable M"  | 
|
1206  | 
shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1207  | 
using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)  | 
| 49774 | 1208  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1209  | 
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
 | 
1210  | 
  fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}"
 | 
| 49774 | 1211  | 
assumes f: "f \<in> borel_measurable M"  | 
1212  | 
assumes g: "g \<in> borel_measurable M"  | 
|
1213  | 
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
 | 
1214  | 
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
 | 
1215  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1216  | 
lemma borel_measurable_prod[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
 | 
1217  | 
  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
 | 
1218  | 
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
 | 
1219  | 
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1220  | 
proof cases  | 
| 
51683
 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 
hoelzl 
parents: 
51478 
diff
changeset
 | 
1221  | 
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
 | 
1222  | 
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
 | 
1223  | 
qed simp  | 
| 49774 | 1224  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1225  | 
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
 | 
1226  | 
  fixes g f :: "'a \<Rightarrow> 'b::{second_countable_topology, metric_space}"
 | 
| 49774 | 1227  | 
assumes f: "f \<in> borel_measurable M"  | 
1228  | 
assumes g: "g \<in> borel_measurable M"  | 
|
1229  | 
shows "(\<lambda>x. dist (f x) (g x)) \<in> borel_measurable M"  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1230  | 
using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)  | 
| 62372 | 1231  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1232  | 
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
 | 
1233  | 
  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
 | 
1234  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1235  | 
assumes g: "g \<in> borel_measurable M"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1236  | 
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
 | 
1237  | 
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
 | 
1238  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1239  | 
lemma borel_measurable_uminus_eq [simp]:  | 
| 
66164
 
2d79288b042c
New theorems and much tidying up of the old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
64911 
diff
changeset
 | 
1240  | 
  fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
 | 
| 
 
2d79288b042c
New theorems and much tidying up of the old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
64911 
diff
changeset
 | 
1241  | 
shows "(\<lambda>x. - f x) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1242  | 
by (smt (verit, ccfv_SIG) borel_measurable_uminus equation_minus_iff measurable_cong)  | 
| 
66164
 
2d79288b042c
New theorems and much tidying up of the old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
64911 
diff
changeset
 | 
1243  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1244  | 
lemma affine_borel_measurable_vector:  | 
| 38656 | 1245  | 
fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"  | 
1246  | 
assumes "f \<in> borel_measurable M"  | 
|
1247  | 
shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"  | 
|
1248  | 
proof (rule borel_measurableI)  | 
|
1249  | 
fix S :: "'x set" assume "open S"  | 
|
1250  | 
show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M"  | 
|
1251  | 
proof cases  | 
|
1252  | 
assume "b \<noteq> 0"  | 
|
| 61808 | 1253  | 
with \<open>open S\<close> 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
 | 
1254  | 
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
 | 
1255  | 
by (auto simp: algebra_simps)  | 
| 47694 | 1256  | 
hence "?S \<in> sets borel" by auto  | 
| 38656 | 1257  | 
moreover  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1258  | 
have "\<And>x. \<lbrakk>a + b *\<^sub>R f x \<in> S\<rbrakk> \<Longrightarrow> f x \<in> (\<lambda>x. (x - a) /\<^sub>R b) ` S"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1259  | 
using \<open>b \<noteq> 0\<close> image_iff by fastforce  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1260  | 
with \<open>b \<noteq> 0\<close> have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1261  | 
by auto  | 
| 40859 | 1262  | 
ultimately show ?thesis using assms unfolding in_borel_measurable_borel  | 
| 38656 | 1263  | 
by auto  | 
1264  | 
qed simp  | 
|
1265  | 
qed  | 
|
1266  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1267  | 
lemma borel_measurable_const_scaleR[measurable (raw)]:  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1268  | 
"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
 | 
1269  | 
using affine_borel_measurable_vector[of f M 0 b] by simp  | 
| 38656 | 1270  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1271  | 
lemma borel_measurable_const_add[measurable (raw)]:  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1272  | 
"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
 | 
1273  | 
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
 | 
1274  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1275  | 
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
 | 
1276  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra"  | 
| 49774 | 1277  | 
assumes f: "f \<in> borel_measurable M"  | 
| 35692 | 1278  | 
shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1279  | 
proof -  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1280  | 
  have "countable {0::'b}" "continuous_on (- {0::'b}) inverse"
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1281  | 
by (auto intro!: continuous_on_inverse continuous_on_id)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1282  | 
then show ?thesis  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1283  | 
by (metis borel_measurable_continuous_countable_exceptions f measurable_compose)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1284  | 
qed  | 
| 35692 | 1285  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1286  | 
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
 | 
1287  | 
"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
 | 
1288  | 
    (\<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
 | 
1289  | 
by (simp add: divide_inverse)  | 
| 38656 | 1290  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1291  | 
lemma borel_measurable_abs[measurable (raw)]:  | 
| 50003 | 1292  | 
"f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"  | 
1293  | 
unfolding abs_real_def by simp  | 
|
| 38656 | 1294  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1295  | 
lemma borel_measurable_nth[measurable (raw)]:  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
1296  | 
"(\<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
 | 
1297  | 
by (simp add: cart_eq_inner_axis)  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41025 
diff
changeset
 | 
1298  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1299  | 
lemma convex_measurable:  | 
| 59415 | 1300  | 
fixes A :: "'a :: euclidean_space set"  | 
| 62372 | 1301  | 
shows "X \<in> borel_measurable M \<Longrightarrow> X ` space M \<subseteq> A \<Longrightarrow> open A \<Longrightarrow> convex_on A q \<Longrightarrow>  | 
| 59415 | 1302  | 
(\<lambda>x. q (X x)) \<in> borel_measurable M"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1303  | 
by (rule measurable_compose[where f=X and N="restrict_space borel A"])  | 
| 59415 | 1304  | 
(auto intro!: borel_measurable_continuous_on_restrict convex_on_continuous measurable_restrict_space2)  | 
| 41830 | 1305  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1306  | 
lemma borel_measurable_ln[measurable (raw)]:  | 
| 49774 | 1307  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
60017
 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 
paulson <lp15@cam.ac.uk> 
parents: 
59658 
diff
changeset
 | 
1308  | 
shows "(\<lambda>x. ln (f x :: real)) \<in> borel_measurable M"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1309  | 
  using borel_measurable_continuous_countable_exceptions[of "{0}"] measurable_compose[OF f]
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1310  | 
by (auto intro!: continuous_on_ln continuous_on_id)  | 
| 41830 | 1311  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1312  | 
lemma borel_measurable_log[measurable (raw)]:  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1313  | 
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log (g x) (f x)) \<in> borel_measurable M"  | 
| 49774 | 1314  | 
unfolding log_def by auto  | 
| 41830 | 1315  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1316  | 
lemma borel_measurable_exp[measurable]:  | 
| 58656 | 1317  | 
  "(exp::'a::{real_normed_field,banach}\<Rightarrow>'a) \<in> borel_measurable borel"
 | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1318  | 
by (intro borel_measurable_continuous_onI continuous_at_imp_continuous_on ballI isCont_exp)  | 
| 50419 | 1319  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1320  | 
lemma measurable_real_floor[measurable]:  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1321  | 
"(floor :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"  | 
| 47761 | 1322  | 
proof -  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1323  | 
have "\<And>a x. \<lfloor>x\<rfloor> = a \<longleftrightarrow> (real_of_int a \<le> x \<and> x < real_of_int (a + 1))"  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1324  | 
by (auto intro: floor_eq2)  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1325  | 
then show ?thesis  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1326  | 
by (auto simp: vimage_def measurable_count_space_eq2_countable)  | 
| 47761 | 1327  | 
qed  | 
1328  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1329  | 
lemma measurable_real_ceiling[measurable]:  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1330  | 
"(ceiling :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1331  | 
unfolding ceiling_def[abs_def] by simp  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1332  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1333  | 
lemma borel_measurable_real_floor: "(\<lambda>x::real. real_of_int \<lfloor>x\<rfloor>) \<in> borel_measurable borel"  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1334  | 
by simp  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1335  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1336  | 
lemma borel_measurable_root [measurable]: "root n \<in> borel_measurable borel"  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1337  | 
by (intro borel_measurable_continuous_onI continuous_intros)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57138 
diff
changeset
 | 
1338  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1339  | 
lemma borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel"  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1340  | 
by (intro borel_measurable_continuous_onI continuous_intros)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57138 
diff
changeset
 | 
1341  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1342  | 
lemma borel_measurable_power [measurable (raw)]:  | 
| 59415 | 1343  | 
  fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}"
 | 
1344  | 
assumes f: "f \<in> borel_measurable M"  | 
|
1345  | 
shows "(\<lambda>x. (f x) ^ n) \<in> borel_measurable M"  | 
|
1346  | 
by (intro borel_measurable_continuous_on [OF _ f] continuous_intros)  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57138 
diff
changeset
 | 
1347  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1348  | 
lemma borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel"  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1349  | 
by (intro borel_measurable_continuous_onI continuous_intros)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57138 
diff
changeset
 | 
1350  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1351  | 
lemma borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel"  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1352  | 
by (intro borel_measurable_continuous_onI continuous_intros)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57138 
diff
changeset
 | 
1353  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1354  | 
lemma borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel"  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1355  | 
by (intro borel_measurable_continuous_onI continuous_intros)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57138 
diff
changeset
 | 
1356  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1357  | 
lemma borel_measurable_sin [measurable]: "(sin :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
 | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1358  | 
by (intro borel_measurable_continuous_onI continuous_intros)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57138 
diff
changeset
 | 
1359  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1360  | 
lemma borel_measurable_cos [measurable]: "(cos :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
 | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1361  | 
by (intro borel_measurable_continuous_onI continuous_intros)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57138 
diff
changeset
 | 
1362  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1363  | 
lemma borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1364  | 
by (intro borel_measurable_continuous_onI continuous_intros)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57138 
diff
changeset
 | 
1365  | 
|
| 70136 | 1366  | 
lemma\<^marker>\<open>tag important\<close> borel_measurable_complex_iff:  | 
| 
57259
 
3a448982a74a
add more derivative and continuity rules for complex-values functions
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1367  | 
"f \<in> borel_measurable M \<longleftrightarrow>  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1368  | 
(\<lambda>x. Re (f x)) \<in> borel_measurable M \<and> (\<lambda>x. Im (f x)) \<in> borel_measurable M" (is "?lhs \<longleftrightarrow> ?rhs")  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1369  | 
proof  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1370  | 
show "?lhs \<Longrightarrow> ?rhs"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1371  | 
using borel_measurable_Im borel_measurable_Re measurable_compose by blast  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1372  | 
assume R: ?rhs  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1373  | 
then have "(\<lambda>x. complex_of_real (Re (f x)) + \<i> * complex_of_real (Im (f x))) \<in> borel_measurable M"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1374  | 
by (intro borel_measurable_add) auto  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1375  | 
then show ?lhs  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1376  | 
using complex_eq by force  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1377  | 
qed  | 
| 
57259
 
3a448982a74a
add more derivative and continuity rules for complex-values functions
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1378  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1379  | 
lemma powr_real_measurable [measurable]:  | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66164 
diff
changeset
 | 
1380  | 
assumes "f \<in> measurable M borel" "g \<in> measurable M borel"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66164 
diff
changeset
 | 
1381  | 
shows "(\<lambda>x. f x powr g x :: real) \<in> measurable M borel"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1382  | 
using assms by (simp_all add: powr_def)  | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66164 
diff
changeset
 | 
1383  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1384  | 
lemma measurable_of_bool[measurable]: "of_bool \<in> count_space UNIV \<rightarrow>\<^sub>M borel"  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1385  | 
by simp  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1386  | 
|
| 69683 | 1387  | 
subsection "Borel space on the extended reals"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1388  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1389  | 
lemma borel_measurable_ereal[measurable (raw)]:  | 
| 43920 | 1390  | 
assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"  | 
| 60771 | 1391  | 
using continuous_on_ereal f by (rule borel_measurable_continuous_on) (rule continuous_on_id)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1392  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1393  | 
lemma borel_measurable_real_of_ereal[measurable (raw)]:  | 
| 62372 | 1394  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
| 49774 | 1395  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1396  | 
shows "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1397  | 
  using measurable_compose[OF f] borel_measurable_continuous_countable_exceptions[of "{\<infinity>, -\<infinity> }"]
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1398  | 
by (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV)  | 
| 49774 | 1399  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1400  | 
lemma borel_measurable_ereal_cases:  | 
| 62372 | 1401  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
| 49774 | 1402  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1403  | 
assumes H: "(\<lambda>x. H (ereal (real_of_ereal (f x)))) \<in> borel_measurable M"  | 
| 49774 | 1404  | 
shows "(\<lambda>x. H (f x)) \<in> borel_measurable M"  | 
1405  | 
proof -  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1406  | 
let ?F = "\<lambda>x. if f x = \<infinity> then H \<infinity> else if f x = - \<infinity> then H (-\<infinity>) else H (ereal (real_of_ereal (f x)))"  | 
| 49774 | 1407  | 
  { 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
 | 
1408  | 
with f H show ?thesis by simp  | 
| 47694 | 1409  | 
qed  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1410  | 
|
| 69739 | 1411  | 
lemma  | 
| 50003 | 1412  | 
fixes f :: "'a \<Rightarrow> ereal" assumes f[measurable]: "f \<in> borel_measurable M"  | 
1413  | 
shows borel_measurable_ereal_abs[measurable(raw)]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"  | 
|
1414  | 
and borel_measurable_ereal_inverse[measurable(raw)]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M"  | 
|
1415  | 
and borel_measurable_uminus_ereal[measurable(raw)]: "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"  | 
|
| 49774 | 1416  | 
by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If)  | 
1417  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1418  | 
lemma borel_measurable_uminus_eq_ereal[simp]:  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1419  | 
"(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1420  | 
by (smt (verit, ccfv_SIG) borel_measurable_uminus_ereal ereal_uminus_uminus measurable_cong)  | 
| 49774 | 1421  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1422  | 
lemma set_Collect_ereal2:  | 
| 62372 | 1423  | 
fixes f g :: "'a \<Rightarrow> ereal"  | 
| 49774 | 1424  | 
assumes f: "f \<in> borel_measurable M"  | 
1425  | 
assumes g: "g \<in> borel_measurable M"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1426  | 
  assumes H: "{x \<in> space M. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))} \<in> sets M"
 | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1427  | 
    "{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
 | 
1428  | 
    "{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
 | 
1429  | 
    "{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
 | 
1430  | 
    "{x \<in> space borel. H (ereal x) (\<infinity>)} \<in> sets borel"
 | 
| 49774 | 1431  | 
  shows "{x \<in> space M. H (f x) (g x)} \<in> sets M"
 | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1432  | 
proof -  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1433  | 
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_of_ereal (g x)))"  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1434  | 
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_of_ereal (f x))) x"  | 
| 49774 | 1435  | 
  { 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
 | 
1436  | 
note * = this  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1437  | 
from assms show ?thesis  | 
| 62390 | 1438  | 
by (subst *) (simp del: space_borel split del: if_split)  | 
| 49774 | 1439  | 
qed  | 
1440  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1441  | 
lemma borel_measurable_ereal_iff:  | 
| 43920 | 1442  | 
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
 | 
1443  | 
proof  | 
| 43920 | 1444  | 
assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"  | 
1445  | 
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
 | 
1446  | 
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
 | 
1447  | 
qed auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1448  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1449  | 
lemma borel_measurable_erealD[measurable_dest]:  | 
| 
59353
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
1450  | 
"(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
1451  | 
unfolding borel_measurable_ereal_iff by simp  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
1452  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1453  | 
theorem borel_measurable_ereal_iff_real:  | 
| 43923 | 1454  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
1455  | 
shows "f \<in> borel_measurable M \<longleftrightarrow>  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1456  | 
    ((\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
 | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1457  | 
proof safe  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1458  | 
  assume *: "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<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
 | 
1459  | 
  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
 | 
1460  | 
  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
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1461  | 
let ?f = "\<lambda>x. if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else ereal (real_of_ereal (f x))"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1462  | 
have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto  | 
| 43920 | 1463  | 
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
 | 
1464  | 
finally show "f \<in> borel_measurable M" .  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1465  | 
qed simp_all  | 
| 41830 | 1466  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1467  | 
lemma borel_measurable_ereal_iff_Iio:  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1468  | 
  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
 | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1469  | 
by (auto simp: borel_Iio measurable_iff_measure_of)  | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1470  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1471  | 
lemma borel_measurable_ereal_iff_Ioi:  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1472  | 
  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
 | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1473  | 
by (auto simp: borel_Ioi measurable_iff_measure_of)  | 
| 35582 | 1474  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1475  | 
lemma vimage_sets_compl_iff:  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1476  | 
"f -` A \<inter> space M \<in> sets M \<longleftrightarrow> f -` (- A) \<inter> space M \<in> sets M"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1477  | 
by (metis Diff_Compl Diff_Diff_Int diff_eq inf_aci(1) sets.Diff sets.top vimage_Compl)  | 
| 49774 | 1478  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1479  | 
lemma borel_measurable_iff_Iic_ereal:  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1480  | 
  "(f::'a\<Rightarrow>ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
 | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1481  | 
  unfolding borel_measurable_ereal_iff_Ioi vimage_sets_compl_iff[where A="{a <..}" for a] by simp
 | 
| 38656 | 1482  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1483  | 
lemma borel_measurable_iff_Ici_ereal:  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1484  | 
  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
 | 
| 
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1485  | 
  unfolding borel_measurable_ereal_iff_Iio vimage_sets_compl_iff[where A="{..< a}" for a] by simp
 | 
| 38656 | 1486  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1487  | 
lemma borel_measurable_ereal2:  | 
| 62372 | 1488  | 
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
 | 
1489  | 
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
 | 
1490  | 
assumes g: "g \<in> borel_measurable M"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1491  | 
assumes H: "(\<lambda>x. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M"  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1492  | 
"(\<lambda>x. H (-\<infinity>) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M"  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1493  | 
"(\<lambda>x. H (\<infinity>) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M"  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1494  | 
"(\<lambda>x. H (ereal (real_of_ereal (f x))) (-\<infinity>)) \<in> borel_measurable M"  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1495  | 
"(\<lambda>x. H (ereal (real_of_ereal (f x))) (\<infinity>)) \<in> borel_measurable M"  | 
| 49774 | 1496  | 
shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1497  | 
proof -  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1498  | 
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_of_ereal (g x)))"  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
1499  | 
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_of_ereal (f x))) x"  | 
| 49774 | 1500  | 
  { 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
 | 
1501  | 
note * = this  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1502  | 
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
 | 
1503  | 
qed  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41969 
diff
changeset
 | 
1504  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1505  | 
lemma [measurable(raw)]:  | 
| 43920 | 1506  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
| 50003 | 1507  | 
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
 | 
1508  | 
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
 | 
1509  | 
and borel_measurable_ereal_times: "(\<lambda>x. f x * g x) \<in> borel_measurable M"  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1510  | 
by (simp_all add: borel_measurable_ereal2)  | 
| 49774 | 1511  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1512  | 
lemma [measurable(raw)]:  | 
| 49774 | 1513  | 
fixes f g :: "'a \<Rightarrow> ereal"  | 
1514  | 
assumes "f \<in> borel_measurable M"  | 
|
1515  | 
assumes "g \<in> borel_measurable M"  | 
|
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1516  | 
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
 | 
1517  | 
and borel_measurable_ereal_divide: "(\<lambda>x. f x / g x) \<in> borel_measurable M"  | 
| 50003 | 1518  | 
using assms by (simp_all add: minus_ereal_def divide_ereal_def)  | 
| 38656 | 1519  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1520  | 
lemma borel_measurable_ereal_sum[measurable (raw)]:  | 
| 43920 | 1521  | 
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"  | 
| 41096 | 1522  | 
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"  | 
1523  | 
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"  | 
|
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1524  | 
using assms by (induction S rule: infinite_finite_induct) auto  | 
| 38656 | 1525  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1526  | 
lemma borel_measurable_ereal_prod[measurable (raw)]:  | 
| 43920 | 1527  | 
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"  | 
| 38656 | 1528  | 
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"  | 
| 41096 | 1529  | 
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"  | 
| 
59361
 
fd5da2434be4
piecewise measurability using restrict_space; cleanup Borel_Space
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
1530  | 
using assms by (induction S rule: infinite_finite_induct) auto  | 
| 38656 | 1531  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1532  | 
lemma borel_measurable_extreal_suminf[measurable (raw)]:  | 
| 43920 | 1533  | 
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"  | 
| 50003 | 1534  | 
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
 | 
1535  | 
shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"  | 
| 50003 | 1536  | 
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp  | 
| 39092 | 1537  | 
|
| 69683 | 1538  | 
subsection "Borel space on the extended non-negative reals"  | 
| 62625 | 1539  | 
|
| 69597 | 1540  | 
text \<open> \<^type>\<open>ennreal\<close> is a topological monoid, so no rules for plus are required, also all order  | 
| 62625 | 1541  | 
statements are usually done on type classes. \<close>  | 
1542  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1543  | 
lemma measurable_enn2ereal[measurable]: "enn2ereal \<in> borel \<rightarrow>\<^sub>M borel"  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1544  | 
by (intro borel_measurable_continuous_onI continuous_on_enn2ereal)  | 
| 62625 | 1545  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1546  | 
lemma measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel"  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1547  | 
by (intro borel_measurable_continuous_onI continuous_on_e2ennreal)  | 
| 62625 | 1548  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1549  | 
lemma borel_measurable_enn2real[measurable (raw)]:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1550  | 
"f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. enn2real (f x)) \<in> M \<rightarrow>\<^sub>M borel"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1551  | 
unfolding enn2real_def[abs_def] by measurable  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1552  | 
|
| 70136 | 1553  | 
definition\<^marker>\<open>tag important\<close> [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"  | 
| 62625 | 1554  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1555  | 
lemma is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) is_borel is_borel"  | 
| 62625 | 1556  | 
unfolding is_borel_def[abs_def]  | 
1557  | 
proof (safe intro!: rel_funI ext dest!: rel_fun_eq_pcr_ennreal[THEN iffD1])  | 
|
1558  | 
fix f and M :: "'a measure"  | 
|
1559  | 
show "f \<in> borel_measurable M" if f: "enn2ereal \<circ> f \<in> borel_measurable M"  | 
|
1560  | 
using measurable_compose[OF f measurable_e2ennreal] by simp  | 
|
1561  | 
qed simp  | 
|
1562  | 
||
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1563  | 
context  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1564  | 
includes ennreal.lifting  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1565  | 
begin  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1566  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1567  | 
lemma measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1568  | 
unfolding is_borel_def[symmetric]  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1569  | 
by transfer simp  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1570  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1571  | 
lemma borel_measurable_ennreal_iff[simp]:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1572  | 
assumes [simp]: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1573  | 
shows "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel \<longleftrightarrow> f \<in> M \<rightarrow>\<^sub>M borel"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1574  | 
proof safe  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1575  | 
assume "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1576  | 
then have "(\<lambda>x. enn2real (ennreal (f x))) \<in> M \<rightarrow>\<^sub>M borel"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1577  | 
by measurable  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1578  | 
then show "f \<in> M \<rightarrow>\<^sub>M borel"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1579  | 
by (rule measurable_cong[THEN iffD1, rotated]) auto  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1580  | 
qed measurable  | 
| 62625 | 1581  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1582  | 
lemma borel_measurable_times_ennreal[measurable (raw)]:  | 
| 62625 | 1583  | 
fixes f g :: "'a \<Rightarrow> ennreal"  | 
1584  | 
shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x * g x) \<in> M \<rightarrow>\<^sub>M borel"  | 
|
1585  | 
unfolding is_borel_def[symmetric] by transfer simp  | 
|
1586  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1587  | 
lemma borel_measurable_inverse_ennreal[measurable (raw)]:  | 
| 62625 | 1588  | 
fixes f :: "'a \<Rightarrow> ennreal"  | 
1589  | 
shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. inverse (f x)) \<in> M \<rightarrow>\<^sub>M borel"  | 
|
1590  | 
unfolding is_borel_def[symmetric] by transfer simp  | 
|
1591  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1592  | 
lemma borel_measurable_divide_ennreal[measurable (raw)]:  | 
| 62625 | 1593  | 
fixes f :: "'a \<Rightarrow> ennreal"  | 
1594  | 
shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x / g x) \<in> M \<rightarrow>\<^sub>M borel"  | 
|
1595  | 
unfolding divide_ennreal_def by simp  | 
|
1596  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1597  | 
lemma borel_measurable_minus_ennreal[measurable (raw)]:  | 
| 62625 | 1598  | 
fixes f :: "'a \<Rightarrow> ennreal"  | 
1599  | 
shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x - g x) \<in> M \<rightarrow>\<^sub>M borel"  | 
|
1600  | 
unfolding is_borel_def[symmetric] by transfer simp  | 
|
1601  | 
||
| 
73253
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71633 
diff
changeset
 | 
1602  | 
lemma borel_measurable_power_ennreal [measurable (raw)]:  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71633 
diff
changeset
 | 
1603  | 
fixes f :: "_ \<Rightarrow> ennreal"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71633 
diff
changeset
 | 
1604  | 
assumes f: "f \<in> borel_measurable M"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71633 
diff
changeset
 | 
1605  | 
shows "(\<lambda>x. (f x) ^ n) \<in> borel_measurable M"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71633 
diff
changeset
 | 
1606  | 
by (induction n) (use f in auto)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71633 
diff
changeset
 | 
1607  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1608  | 
lemma borel_measurable_prod_ennreal[measurable (raw)]:  | 
| 62625 | 1609  | 
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ennreal"  | 
1610  | 
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"  | 
|
1611  | 
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1612  | 
using assms by (induction S rule: infinite_finite_induct) auto  | 
| 62625 | 1613  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1614  | 
end  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62625 
diff
changeset
 | 
1615  | 
|
| 62625 | 1616  | 
hide_const (open) is_borel  | 
1617  | 
||
| 69683 | 1618  | 
subsection \<open>LIMSEQ is borel measurable\<close>  | 
| 39092 | 1619  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1620  | 
lemma borel_measurable_LIMSEQ_real:  | 
| 39092 | 1621  | 
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"  | 
| 61969 | 1622  | 
assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"  | 
| 39092 | 1623  | 
and u: "\<And>i. u i \<in> borel_measurable M"  | 
1624  | 
shows "u' \<in> borel_measurable M"  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1625  | 
proof -  | 
| 43920 | 1626  | 
have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)"  | 
| 46731 | 1627  | 
using u' by (simp add: lim_imp_Liminf)  | 
| 43920 | 1628  | 
moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M"  | 
| 39092 | 1629  | 
by auto  | 
| 43920 | 1630  | 
ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff)  | 
| 39092 | 1631  | 
qed  | 
1632  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1633  | 
lemma borel_measurable_LIMSEQ_metric:  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1634  | 
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
 | 
1635  | 
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"  | 
| 61969 | 1636  | 
assumes lim: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. f i x) \<longlonglongrightarrow> g x"  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1637  | 
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
 | 
1638  | 
unfolding borel_eq_closed  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1639  | 
proof (safe intro!: measurable_measure_of)  | 
| 62372 | 1640  | 
fix A :: "'b set" assume "closed A"  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1641  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1642  | 
have [measurable]: "(\<lambda>x. infdist (g x) A) \<in> borel_measurable M"  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1643  | 
proof (rule borel_measurable_LIMSEQ_real)  | 
| 61969 | 1644  | 
show "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. infdist (f i x) A) \<longlonglongrightarrow> infdist (g x) A"  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1645  | 
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
 | 
1646  | 
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
 | 
1647  | 
by (intro borel_measurable_continuous_on[where f="\<lambda>x. infdist x A"]  | 
| 
60150
 
bd773c47ad0b
New material about complex transcendental functions (especially Ln, Arg) and polynomials
 
paulson <lp15@cam.ac.uk> 
parents: 
60017 
diff
changeset
 | 
1648  | 
continuous_at_imp_continuous_on ballI continuous_infdist continuous_ident) auto  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1649  | 
qed  | 
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1650  | 
|
| 
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1651  | 
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
 | 
1652  | 
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
 | 
1653  | 
    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
 | 
1654  | 
then have "\<And>x. infdist x A = 0 \<longleftrightarrow> x \<in> A"  | 
| 61808 | 1655  | 
using \<open>closed A\<close> by (simp add: in_closed_iff_infdist_zero)  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1656  | 
    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
 | 
1657  | 
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
 | 
1658  | 
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
 | 
1659  | 
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
 | 
1660  | 
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
 | 
1661  | 
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
 | 
1662  | 
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
 | 
1663  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1664  | 
lemma sets_Collect_Cauchy[measurable]:  | 
| 57036 | 1665  | 
  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
 | 
1666  | 
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"  | 
| 49774 | 1667  | 
  shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M"
 | 
| 57036 | 1668  | 
unfolding metric_Cauchy_iff2 using f by auto  | 
| 49774 | 1669  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1670  | 
lemma borel_measurable_lim_metric[measurable (raw)]:  | 
| 57036 | 1671  | 
  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
 | 
1672  | 
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"  | 
| 49774 | 1673  | 
shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"  | 
1674  | 
proof -  | 
|
| 63040 | 1675  | 
define u' where "u' x = lim (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)" for x  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1676  | 
then have *: "\<And>x. lim (\<lambda>i. f i x) = (if Cauchy (\<lambda>i. f i x) then u' x else (THE x. False))"  | 
| 64287 | 1677  | 
by (auto simp: lim_def convergent_eq_Cauchy[symmetric])  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1678  | 
have "u' \<in> borel_measurable M"  | 
| 57036 | 1679  | 
proof (rule borel_measurable_LIMSEQ_metric)  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1680  | 
fix x  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1681  | 
have "convergent (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)"  | 
| 49774 | 1682  | 
by (cases "Cauchy (\<lambda>i. f i x)")  | 
| 64287 | 1683  | 
(auto simp add: convergent_eq_Cauchy[symmetric] convergent_def)  | 
| 61969 | 1684  | 
then show "(\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0) \<longlonglongrightarrow> u' x"  | 
| 62372 | 1685  | 
unfolding u'_def  | 
| 
50002
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1686  | 
by (rule convergent_LIMSEQ_iff[THEN iffD1])  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1687  | 
qed measurable  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1688  | 
then show ?thesis  | 
| 
 
ce0d316b5b44
add measurability prover; add support for Borel sets
 
hoelzl 
parents: 
50001 
diff
changeset
 | 
1689  | 
unfolding * by measurable  | 
| 49774 | 1690  | 
qed  | 
1691  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1692  | 
lemma borel_measurable_suminf[measurable (raw)]:  | 
| 57036 | 1693  | 
  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
 | 
1694  | 
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"  | 
| 49774 | 1695  | 
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
 | 
1696  | 
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp  | 
| 49774 | 1697  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1698  | 
lemma Collect_closed_imp_pred_borel: "closed {x. P x} \<Longrightarrow> Measurable.pred borel P"
 | 
| 63389 | 1699  | 
by (simp add: pred_def)  | 
1700  | 
||
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1701  | 
text \<open>Proof by Jeremy Avigad and Luke Serafin\<close>  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1702  | 
lemma isCont_borel_pred[measurable]:  | 
| 63389 | 1703  | 
fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"  | 
1704  | 
shows "Measurable.pred borel (isCont f)"  | 
|
1705  | 
proof (subst measurable_cong)  | 
|
1706  | 
let ?I = "\<lambda>j. inverse(real (Suc j))"  | 
|
1707  | 
show "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)" for x  | 
|
1708  | 
unfolding continuous_at_eps_delta  | 
|
1709  | 
proof safe  | 
|
1710  | 
fix i assume "\<forall>e>0. \<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e"  | 
|
1711  | 
moreover have "0 < ?I i / 2"  | 
|
1712  | 
by simp  | 
|
1713  | 
ultimately obtain d where d: "0 < d" "\<And>y. dist x y < d \<Longrightarrow> dist (f y) (f x) < ?I i / 2"  | 
|
1714  | 
by (metis dist_commute)  | 
|
1715  | 
then obtain j where j: "?I j < d"  | 
|
1716  | 
by (metis reals_Archimedean)  | 
|
1717  | 
||
1718  | 
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"  | 
|
1719  | 
proof (safe intro!: exI[where x=j])  | 
|
1720  | 
fix y z assume *: "dist x y < ?I j" "dist x z < ?I j"  | 
|
1721  | 
have "dist (f y) (f z) \<le> dist (f y) (f x) + dist (f z) (f x)"  | 
|
1722  | 
by (rule dist_triangle2)  | 
|
1723  | 
also have "\<dots> < ?I i / 2 + ?I i / 2"  | 
|
1724  | 
by (intro add_strict_mono d less_trans[OF _ j] *)  | 
|
1725  | 
also have "\<dots> \<le> ?I i"  | 
|
| 71633 | 1726  | 
by (simp add: field_simps)  | 
| 63389 | 1727  | 
finally show "dist (f y) (f z) \<le> ?I i"  | 
1728  | 
by simp  | 
|
1729  | 
qed  | 
|
1730  | 
next  | 
|
1731  | 
fix e::real assume "0 < e"  | 
|
1732  | 
then obtain n where n: "?I n < e"  | 
|
1733  | 
by (metis reals_Archimedean)  | 
|
1734  | 
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"  | 
|
1735  | 
from this[THEN spec, of "Suc n"]  | 
|
1736  | 
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)"  | 
|
1737  | 
by auto  | 
|
1738  | 
||
1739  | 
show "\<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e"  | 
|
1740  | 
proof (safe intro!: exI[of _ "?I j"])  | 
|
1741  | 
fix y assume "dist y x < ?I j"  | 
|
1742  | 
then have "dist (f y) (f x) \<le> ?I (Suc n)"  | 
|
1743  | 
by (intro j) (auto simp: dist_commute)  | 
|
1744  | 
also have "?I (Suc n) < ?I n"  | 
|
1745  | 
by simp  | 
|
1746  | 
also note n  | 
|
1747  | 
finally show "dist (f y) (f x) < e" .  | 
|
1748  | 
qed simp  | 
|
1749  | 
qed  | 
|
1750  | 
qed (intro pred_intros_countable closed_Collect_all closed_Collect_le open_Collect_less  | 
|
1751  | 
Collect_closed_imp_pred_borel closed_Collect_imp open_Collect_conj continuous_intros)  | 
|
1752  | 
||
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1753  | 
lemma isCont_borel:  | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57275 
diff
changeset
 | 
1754  | 
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
 | 
1755  | 
  shows "{x. isCont f x} \<in> sets borel"
 | 
| 63389 | 1756  | 
by simp  | 
| 62083 | 1757  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1758  | 
lemma is_real_interval:  | 
| 
61880
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1759  | 
assumes S: "is_interval S"  | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1760  | 
  shows "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or> S = {a<..} \<or> S = {a..} \<or>
 | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1761  | 
    S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}"
 | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1762  | 
using S unfolding is_interval_1 by (blast intro: interval_cases)  | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1763  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1764  | 
lemma real_interval_borel_measurable:  | 
| 
61880
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1765  | 
assumes "is_interval (S::real set)"  | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1766  | 
shows "S \<in> sets borel"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1767  | 
proof -  | 
| 
61880
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1768  | 
  from assms is_real_interval have "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or>
 | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1769  | 
    S = {a<..} \<or> S = {a..} \<or> S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}" by auto
 | 
| 74362 | 1770  | 
then show ?thesis  | 
| 
61880
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1771  | 
by auto  | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1772  | 
qed  | 
| 
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1773  | 
|
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
1774  | 
text \<open>The next lemmas hold in any second countable linorder (including ennreal or ereal for instance),  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
1775  | 
but in the current state they are restricted to reals.\<close>  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
1776  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1777  | 
lemma borel_measurable_mono_on_fnc:  | 
| 62083 | 1778  | 
fixes f :: "real \<Rightarrow> real" and A :: "real set"  | 
| 
75607
 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 
desharna 
parents: 
74362 
diff
changeset
 | 
1779  | 
assumes "mono_on A f"  | 
| 62083 | 1780  | 
shows "f \<in> borel_measurable (restrict_space borel A)"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1781  | 
proof -  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1782  | 
  have "\<And>x. x \<in> A \<Longrightarrow> {x} \<in> sets (restrict_space borel A)"
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1783  | 
using sets_restrict_space by fastforce  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1784  | 
moreover  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1785  | 
  have "continuous_on (A \<inter> - {a \<in> A. \<not> continuous (at a within A) f}) f"
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1786  | 
by (force simp: continuous_on_eq_continuous_within intro: continuous_within_subset)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1787  | 
then have "f \<in> borel_measurable (restrict_space (restrict_space borel A)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1788  | 
              (- {a \<in> A. \<not> continuous (at a within A) f}))"
 | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1789  | 
by (smt (verit, best) borel_measurable_continuous_on_restrict measurable_cong_sets sets_restrict_restrict_space)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1790  | 
ultimately show ?thesis  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1791  | 
using measurable_restrict_countable[OF mono_on_ctble_discont[OF assms]]  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1792  | 
by (smt (verit, del_insts) UNIV_I mem_Collect_eq space_borel)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1793  | 
qed  | 
| 62083 | 1794  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1795  | 
lemma borel_measurable_piecewise_mono:  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
1796  | 
fixes f::"real \<Rightarrow> real" and C::"real set set"  | 
| 
75607
 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 
desharna 
parents: 
74362 
diff
changeset
 | 
1797  | 
assumes "countable C" "\<And>c. c \<in> C \<Longrightarrow> c \<in> sets borel" "\<And>c. c \<in> C \<Longrightarrow> mono_on c f" "(\<Union>C) = UNIV"  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
1798  | 
shows "f \<in> borel_measurable borel"  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68635 
diff
changeset
 | 
1799  | 
by (rule measurable_piecewise_restrict[of C], auto intro: borel_measurable_mono_on_fnc simp: assms)  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
1800  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1801  | 
lemma borel_measurable_mono:  | 
| 
61880
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1802  | 
fixes f :: "real \<Rightarrow> real"  | 
| 62083 | 1803  | 
shows "mono f \<Longrightarrow> f \<in> borel_measurable borel"  | 
| 
75607
 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 
desharna 
parents: 
74362 
diff
changeset
 | 
1804  | 
using borel_measurable_mono_on_fnc[of UNIV f] by (simp add: mono_def mono_on_def)  | 
| 
61880
 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 
hoelzl 
parents: 
61808 
diff
changeset
 | 
1805  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1806  | 
lemma measurable_bdd_below_real[measurable (raw)]:  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1807  | 
fixes F :: "'a \<Rightarrow> 'i \<Rightarrow> real"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1808  | 
assumes [simp]: "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> M \<rightarrow>\<^sub>M borel"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1809  | 
shows "Measurable.pred M (\<lambda>x. bdd_below ((\<lambda>i. F i x)`I))"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1810  | 
proof (subst measurable_cong)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1811  | 
show "bdd_below ((\<lambda>i. F i x)`I) \<longleftrightarrow> (\<exists>q\<in>\<int>. \<forall>i\<in>I. q \<le> F i x)" for x  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1812  | 
by (auto simp: bdd_below_def intro!: bexI[of _ "of_int (floor _)"] intro: order_trans of_int_floor_le)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1813  | 
show "Measurable.pred M (\<lambda>w. \<exists>q\<in>\<int>. \<forall>i\<in>I. q \<le> F i w)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1814  | 
using countable_int by measurable  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1815  | 
qed  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1816  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1817  | 
lemma borel_measurable_cINF_real[measurable (raw)]:  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1818  | 
fixes F :: "_ \<Rightarrow> _ \<Rightarrow> real"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1819  | 
assumes [simp]: "countable I"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1820  | 
assumes F[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
1821  | 
shows "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable M"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1822  | 
proof (rule measurable_piecewise_restrict)  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1823  | 
  let ?\<Omega> = "{x\<in>space M. bdd_below ((\<lambda>i. F i x)`I)}"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1824  | 
  show "countable {?\<Omega>, - ?\<Omega>}" "space M \<subseteq> \<Union>{?\<Omega>, - ?\<Omega>}" "\<And>X. X \<in> {?\<Omega>, - ?\<Omega>} \<Longrightarrow> X \<inter> space M \<in> sets M"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1825  | 
by auto  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
1826  | 
  fix X assume "X \<in> {?\<Omega>, - ?\<Omega>}" then show "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable (restrict_space M X)"
 | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1827  | 
proof safe  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
1828  | 
show "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable (restrict_space M ?\<Omega>)"  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1829  | 
by (intro borel_measurable_cINF measurable_restrict_space1 F)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1830  | 
(auto simp: space_restrict_space)  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
1831  | 
show "(\<lambda>x. INF i\<in>I. F i x) \<in> borel_measurable (restrict_space M (-?\<Omega>))"  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1832  | 
proof (subst measurable_cong)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1833  | 
fix x assume "x \<in> space (restrict_space M (-?\<Omega>))"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1834  | 
then have "\<not> (\<forall>i\<in>I. - F i x \<le> y)" for y  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1835  | 
by (auto simp: space_restrict_space bdd_above_def bdd_above_uminus[symmetric])  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69022 
diff
changeset
 | 
1836  | 
then show "(INF i\<in>I. F i x) = - (THE x. False)"  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1837  | 
by (auto simp: space_restrict_space Inf_real_def Sup_real_def Least_def simp del: Set.ball_simps(10))  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1838  | 
qed simp  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1839  | 
qed  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1840  | 
qed  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1841  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1842  | 
lemma borel_Ici: "borel = sigma UNIV (range (\<lambda>x::real. {x ..}))"
 | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1843  | 
proof (safe intro!: borel_eq_sigmaI1[OF borel_Iio])  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1844  | 
fix x :: real  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1845  | 
  have eq: "{..<x} = space (sigma UNIV (range atLeast)) - {x ..}"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1846  | 
by auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1847  | 
  show "{..<x} \<in> sets (sigma UNIV (range atLeast))"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1848  | 
unfolding eq by (intro sets.compl_sets) auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1849  | 
qed auto  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1850  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1851  | 
lemma borel_measurable_pred_less[measurable (raw)]:  | 
| 
64008
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1852  | 
  fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
 | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1853  | 
shows "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> Measurable.pred M (\<lambda>w. f w < g w)"  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1854  | 
unfolding Measurable.pred_def by (rule borel_measurable_less)  | 
| 
 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 
hoelzl 
parents: 
63952 
diff
changeset
 | 
1855  | 
|
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
1856  | 
no_notation  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
1857  | 
eucl_less (infix "<e" 50)  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54230 
diff
changeset
 | 
1858  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1859  | 
lemma borel_measurable_Max2[measurable (raw)]:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1860  | 
  fixes f::"_ \<Rightarrow> _ \<Rightarrow> 'a::{second_countable_topology, dense_linorder, linorder_topology}"
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1861  | 
assumes "finite I"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1862  | 
and [measurable]: "\<And>i. f i \<in> borel_measurable M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1863  | 
  shows "(\<lambda>x. Max{f i x |i. i \<in> I}) \<in> borel_measurable M"
 | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1864  | 
by (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image)  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1865  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1866  | 
lemma measurable_compose_n [measurable (raw)]:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1867  | 
assumes "T \<in> measurable M M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1868  | 
shows "(T^^n) \<in> measurable M M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1869  | 
by (induction n, auto simp add: measurable_compose[OF _ assms])  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1870  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1871  | 
lemma measurable_real_imp_nat:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1872  | 
fixes f::"'a \<Rightarrow> nat"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1873  | 
assumes [measurable]: "(\<lambda>x. real(f x)) \<in> borel_measurable M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1874  | 
shows "f \<in> measurable M (count_space UNIV)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1875  | 
proof -  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1876  | 
let ?g = "(\<lambda>x. real(f x))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1877  | 
  have "\<And>(n::nat). ?g-`({real n}) \<inter> space M = f-`{n} \<inter> space M" by auto
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1878  | 
  moreover have "\<And>(n::nat). ?g-`({real n}) \<inter> space M \<in> sets M" using assms by measurable
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1879  | 
  ultimately have "\<And>(n::nat). f-`{n} \<inter> space M \<in> sets M" by simp
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1880  | 
then show ?thesis using measurable_count_space_eq2_countable by blast  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1881  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1882  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1883  | 
lemma measurable_equality_set [measurable]:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1884  | 
  fixes f g::"_\<Rightarrow> 'a::{second_countable_topology, t2_space}"
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1885  | 
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1886  | 
  shows "{x \<in> space M. f x = g x} \<in> sets M"
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1887  | 
proof -  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1888  | 
  define A where "A = {x \<in> space M. f x = g x}"
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1889  | 
  define B where "B = {y. \<exists>x::'a. y = (x,x)}"
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1890  | 
have "A = (\<lambda>x. (f x, g x))-`B \<inter> space M" unfolding A_def B_def by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1891  | 
moreover have "(\<lambda>x. (f x, g x)) \<in> borel_measurable M" by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1892  | 
moreover have "B \<in> sets borel" unfolding B_def by (simp add: closed_diagonal)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1893  | 
ultimately have "A \<in> sets M" by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1894  | 
then show ?thesis unfolding A_def by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1895  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1896  | 
|
| 
78519
 
f675e2a31682
A subtle fix involving the "measurable" attribute
 
paulson <lp15@cam.ac.uk> 
parents: 
78516 
diff
changeset
 | 
1897  | 
text \<open>Logically equivalent to those with the opposite orientation, still these are needed\<close>  | 
| 
 
f675e2a31682
A subtle fix involving the "measurable" attribute
 
paulson <lp15@cam.ac.uk> 
parents: 
78516 
diff
changeset
 | 
1898  | 
lemma measurable_inequality_set_flipped:  | 
| 
 
f675e2a31682
A subtle fix involving the "measurable" attribute
 
paulson <lp15@cam.ac.uk> 
parents: 
78516 
diff
changeset
 | 
1899  | 
  fixes f g::"_ \<Rightarrow> 'a::{second_countable_topology, linorder_topology}"
 | 
| 
 
f675e2a31682
A subtle fix involving the "measurable" attribute
 
paulson <lp15@cam.ac.uk> 
parents: 
78516 
diff
changeset
 | 
1900  | 
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"  | 
| 
 
f675e2a31682
A subtle fix involving the "measurable" attribute
 
paulson <lp15@cam.ac.uk> 
parents: 
78516 
diff
changeset
 | 
1901  | 
  shows "{x \<in> space M. f x \<ge> g x} \<in> sets M"
 | 
| 
 
f675e2a31682
A subtle fix involving the "measurable" attribute
 
paulson <lp15@cam.ac.uk> 
parents: 
78516 
diff
changeset
 | 
1902  | 
        "{x \<in> space M. f x > g x} \<in> sets M"
 | 
| 
 
f675e2a31682
A subtle fix involving the "measurable" attribute
 
paulson <lp15@cam.ac.uk> 
parents: 
78516 
diff
changeset
 | 
1903  | 
by auto  | 
| 
 
f675e2a31682
A subtle fix involving the "measurable" attribute
 
paulson <lp15@cam.ac.uk> 
parents: 
78516 
diff
changeset
 | 
1904  | 
|
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
1905  | 
lemmas measurable_inequality_set [measurable] =  | 
| 
78519
 
f675e2a31682
A subtle fix involving the "measurable" attribute
 
paulson <lp15@cam.ac.uk> 
parents: 
78516 
diff
changeset
 | 
1906  | 
borel_measurable_le borel_measurable_less measurable_inequality_set_flipped  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1907  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1908  | 
proposition measurable_limit [measurable]:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1909  | 
fixes f::"nat \<Rightarrow> 'a \<Rightarrow> 'b::first_countable_topology"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1910  | 
assumes [measurable]: "\<And>n::nat. f n \<in> borel_measurable M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1911  | 
shows "Measurable.pred M (\<lambda>x. (\<lambda>n. f n x) \<longlonglongrightarrow> c)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1912  | 
proof -  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1913  | 
obtain A :: "nat \<Rightarrow> 'b set" where A:  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1914  | 
"\<And>i. open (A i)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1915  | 
"\<And>i. c \<in> A i"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1916  | 
"\<And>S. open S \<Longrightarrow> c \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1917  | 
by (rule countable_basis_at_decseq) blast  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1918  | 
|
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1919  | 
have [measurable]: "\<And>N i. (f N)-`(A i) \<inter> space M \<in> sets M" using A(1) by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1920  | 
  then have mes: "(\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i) \<inter> space M) \<in> sets M" by blast
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1921  | 
|
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1922  | 
have "(u \<longlonglongrightarrow> c) \<longleftrightarrow> (\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)" for u::"nat \<Rightarrow> 'b"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1923  | 
proof  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1924  | 
assume "u \<longlonglongrightarrow> c"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1925  | 
then have "eventually (\<lambda>n. u n \<in> A i) sequentially" for i using A(1)[of i] A(2)[of i]  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1926  | 
by (simp add: topological_tendstoD)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1927  | 
then show "(\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)" by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1928  | 
next  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1929  | 
assume H: "(\<forall>i. eventually (\<lambda>n. u n \<in> A i) sequentially)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1930  | 
show "(u \<longlonglongrightarrow> c)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1931  | 
proof (rule topological_tendstoI)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1932  | 
fix S assume "open S" "c \<in> S"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1933  | 
with A(3)[OF this] obtain i where "A i \<subseteq> S"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1934  | 
using eventually_False_sequentially eventually_mono by blast  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1935  | 
moreover have "eventually (\<lambda>n. u n \<in> A i) sequentially" using H by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1936  | 
ultimately show "\<forall>\<^sub>F n in sequentially. u n \<in> S"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1937  | 
by (simp add: eventually_mono subset_eq)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1938  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1939  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1940  | 
  then have "{x. (\<lambda>n. f n x) \<longlonglongrightarrow> c} = (\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i))"
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1941  | 
by (auto simp add: atLeast_def eventually_at_top_linorder)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1942  | 
  then have "{x \<in> space M. (\<lambda>n. f n x) \<longlonglongrightarrow> c} = (\<Inter>i. \<Union>n. \<Inter>N\<in>{n..}. (f N)-`(A i) \<inter> space M)"
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1943  | 
by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1944  | 
  then have "{x \<in> space M. (\<lambda>n. f n x) \<longlonglongrightarrow> c} \<in> sets M" using mes by simp
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1945  | 
then show ?thesis by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1946  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1947  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1948  | 
lemma measurable_limit2 [measurable]:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1949  | 
fixes u::"nat \<Rightarrow> 'a \<Rightarrow> real"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1950  | 
assumes [measurable]: "\<And>n. u n \<in> borel_measurable M" "v \<in> borel_measurable M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1951  | 
shows "Measurable.pred M (\<lambda>x. (\<lambda>n. u n x) \<longlonglongrightarrow> v x)"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1952  | 
proof -  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1953  | 
define w where "w = (\<lambda>n x. u n x - v x)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1954  | 
have [measurable]: "w n \<in> borel_measurable M" for n unfolding w_def by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1955  | 
have "((\<lambda>n. u n x) \<longlonglongrightarrow> v x) \<longleftrightarrow> ((\<lambda>n. w n x) \<longlonglongrightarrow> 0)" for x  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1956  | 
unfolding w_def using Lim_null by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1957  | 
then show ?thesis using measurable_limit by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1958  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1959  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1960  | 
lemma measurable_P_restriction [measurable (raw)]:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1961  | 
assumes [measurable]: "Measurable.pred M P" "A \<in> sets M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1962  | 
  shows "{x \<in> A. P x} \<in> sets M"
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1963  | 
proof -  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1964  | 
have "A \<subseteq> space M" using sets.sets_into_space[OF assms(2)].  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1965  | 
  then have "{x \<in> A. P x} = A \<inter> {x \<in> space M. P x}" by blast
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1966  | 
then show ?thesis by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1967  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1968  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1969  | 
lemma measurable_sum_nat [measurable (raw)]:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1970  | 
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> nat"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1971  | 
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> measurable M (count_space UNIV)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1972  | 
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> measurable M (count_space UNIV)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1973  | 
proof cases  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1974  | 
assume "finite S"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1975  | 
then show ?thesis using assms by induct auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1976  | 
qed simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1977  | 
|
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1978  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1979  | 
lemma measurable_abs_powr [measurable]:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1980  | 
fixes p::real  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1981  | 
assumes [measurable]: "f \<in> borel_measurable M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1982  | 
shows "(\<lambda>x. \<bar>f x\<bar> powr p) \<in> borel_measurable M"  | 
| 
70688
 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 
paulson <lp15@cam.ac.uk> 
parents: 
70617 
diff
changeset
 | 
1983  | 
by simp  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1984  | 
|
| 69566 | 1985  | 
text \<open>The next one is a variation around \<open>measurable_restrict_space\<close>.\<close>  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1986  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1987  | 
lemma measurable_restrict_space3:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1988  | 
assumes "f \<in> measurable M N" and  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1989  | 
"f \<in> A \<rightarrow> B"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1990  | 
shows "f \<in> measurable (restrict_space M A) (restrict_space N B)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1991  | 
proof -  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1992  | 
have "f \<in> measurable (restrict_space M A) N" using assms(1) measurable_restrict_space1 by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1993  | 
then show ?thesis by (metis Int_iff funcsetI funcset_mem  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1994  | 
measurable_restrict_space2[of f, of "restrict_space M A", of B, of N] assms(2) space_restrict_space)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1995  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
1996  | 
|
| 
70688
 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 
paulson <lp15@cam.ac.uk> 
parents: 
70617 
diff
changeset
 | 
1997  | 
lemma measurable_restrict_mono:  | 
| 
 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 
paulson <lp15@cam.ac.uk> 
parents: 
70617 
diff
changeset
 | 
1998  | 
assumes f: "f \<in> restrict_space M A \<rightarrow>\<^sub>M N" and "B \<subseteq> A"  | 
| 
 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 
paulson <lp15@cam.ac.uk> 
parents: 
70617 
diff
changeset
 | 
1999  | 
shows "f \<in> restrict_space M B \<rightarrow>\<^sub>M N"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2000  | 
by (rule measurable_compose[OF measurable_restrict_space3 f]) (use \<open>B \<subseteq> A\<close> in auto)  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2001  | 
|
| 
70688
 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 
paulson <lp15@cam.ac.uk> 
parents: 
70617 
diff
changeset
 | 
2002  | 
|
| 69566 | 2003  | 
text \<open>The next one is a variation around \<open>measurable_piecewise_restrict\<close>.\<close>  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2004  | 
|
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
2005  | 
lemma measurable_piecewise_restrict2:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2006  | 
assumes [measurable]: "\<And>n. A n \<in> sets M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2007  | 
and "space M = (\<Union>(n::nat). A n)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2008  | 
"\<And>n. \<exists>h \<in> measurable M N. (\<forall>x \<in> A n. f x = h x)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2009  | 
shows "f \<in> measurable M N"  | 
| 
69652
 
3417a8f154eb
updated tagging first 5
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
2010  | 
proof (rule measurableI)  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2011  | 
fix B assume [measurable]: "B \<in> sets N"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2012  | 
  {
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2013  | 
fix n::nat  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2014  | 
obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2015  | 
using assms(3) by blast  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2016  | 
then have *: "f-`B \<inter> A n = h-`B \<inter> A n" by auto  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2017  | 
have "h-`B \<inter> A n = h-`B \<inter> space M \<inter> A n"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2018  | 
using assms(2) sets.sets_into_space by auto  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2019  | 
then have "f-`B \<inter> A n \<in> sets M"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2020  | 
by (simp add: "*")  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2021  | 
}  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2022  | 
then have "(\<Union>n. f-`B \<inter> A n) \<in> sets M"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2023  | 
by measurable  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2024  | 
moreover have "f-`B \<inter> space M = (\<Union>n. f-`B \<inter> A n)"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2025  | 
using assms(2) by blast  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2026  | 
ultimately show "f-`B \<inter> space M \<in> sets M" by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2027  | 
next  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2028  | 
fix x assume "x \<in> space M"  | 
| 
78516
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2029  | 
then obtain n where "x \<in> A n"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2030  | 
using assms(2) by blast  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2031  | 
obtain h where [measurable]: "h \<in> measurable M N" and "\<forall>x \<in> A n. f x = h x"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2032  | 
using assms(3) by blast  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2033  | 
then show "f x \<in> space N"  | 
| 
 
56a408fa2440
substantial tidy-up, shortening many proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
75607 
diff
changeset
 | 
2034  | 
by (metis \<open>x \<in> A n\<close> \<open>x \<in> space M\<close> measurable_space)  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2035  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
2036  | 
|
| 
51683
 
baefa3b461c2
generalize Borel-set properties from real/ereal/ordered_euclidean_spaces to order_topology and real_normed_vector
 
hoelzl 
parents: 
51478 
diff
changeset
 | 
2037  | 
end  |