author | paulson |
Tue, 12 Jan 2010 16:55:59 +0000 | |
changeset 34883 | 77f0d11dec76 |
parent 33657 | a4179bf442d1 |
child 35028 | 108662d50512 |
permissions | -rw-r--r-- |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
1 |
header {*Borel Sets*} |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
2 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
3 |
theory Borel |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
4 |
imports Measure |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
5 |
begin |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
6 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
7 |
text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*} |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
8 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
9 |
definition borel_space where |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
10 |
"borel_space = sigma (UNIV::real set) (range (\<lambda>a::real. {x. x \<le> a}))" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
11 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
12 |
definition borel_measurable where |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
13 |
"borel_measurable a = measurable a borel_space" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
14 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
15 |
definition indicator_fn where |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
16 |
"indicator_fn s = (\<lambda>x. if x \<in> s then 1 else (0::real))" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
17 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
18 |
definition mono_convergent where |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
19 |
"mono_convergent u f s \<equiv> |
33536 | 20 |
(\<forall>x m n. m \<le> n \<and> x \<in> s \<longrightarrow> u m x \<le> u n x) \<and> |
21 |
(\<forall>x \<in> s. (\<lambda>i. u i x) ----> f x)" |
|
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
22 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
23 |
lemma in_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
24 |
"f \<in> borel_measurable M \<longleftrightarrow> |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
25 |
sigma_algebra M \<and> |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
26 |
(\<forall>s \<in> sets (sigma UNIV (range (\<lambda>a::real. {x. x \<le> a}))). |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
27 |
f -` s \<inter> space M \<in> sets M)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
28 |
apply (auto simp add: borel_measurable_def measurable_def borel_space_def) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
29 |
apply (metis PowD UNIV_I Un_commute sigma_algebra_sigma subset_Pow_Union subset_UNIV subset_Un_eq) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
30 |
done |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
31 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
32 |
lemma (in sigma_algebra) borel_measurable_const: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
33 |
"(\<lambda>x. c) \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
34 |
by (auto simp add: in_borel_measurable prems) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
35 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
36 |
lemma (in sigma_algebra) borel_measurable_indicator: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
37 |
assumes a: "a \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
38 |
shows "indicator_fn a \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
39 |
apply (auto simp add: in_borel_measurable indicator_fn_def prems) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
40 |
apply (metis Diff_eq Int_commute a compl_sets) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
41 |
done |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
42 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
43 |
lemma Collect_eq: "{w \<in> X. f w \<le> a} = {w. f w \<le> a} \<inter> X" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
44 |
by (metis Collect_conj_eq Collect_mem_eq Int_commute) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
45 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
46 |
lemma (in measure_space) borel_measurable_le_iff: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
47 |
"f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
48 |
proof |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
49 |
assume f: "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
50 |
{ fix a |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
51 |
have "{w \<in> space M. f w \<le> a} \<in> sets M" using f |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
52 |
apply (auto simp add: in_borel_measurable sigma_def Collect_eq) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
53 |
apply (drule_tac x="{x. x \<le> a}" in bspec, auto) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
54 |
apply (metis equalityE rangeI subsetD sigma_sets.Basic) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
55 |
done |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
56 |
} |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
57 |
thus "\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M" by blast |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
58 |
next |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
59 |
assume "\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
60 |
thus "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
61 |
apply (simp add: borel_measurable_def borel_space_def Collect_eq) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
62 |
apply (rule measurable_sigma, auto) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
63 |
done |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
64 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
65 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
66 |
lemma Collect_less_le: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
67 |
"{w \<in> X. f w < g w} = (\<Union>n. {w \<in> X. f w \<le> g w - inverse(real(Suc n))})" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
68 |
proof auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
69 |
fix w |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
70 |
assume w: "f w < g w" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
71 |
hence nz: "g w - f w \<noteq> 0" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
72 |
by arith |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
73 |
with w have "real(Suc(natceiling(inverse(g w - f w)))) > inverse(g w - f w)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
74 |
by (metis lessI order_le_less_trans real_natceiling_ge real_of_nat_less_iff) hence "inverse(real(Suc(natceiling(inverse(g w - f w))))) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
75 |
< inverse(inverse(g w - f w))" |
33657 | 76 |
by (metis less_iff_diff_less_0 less_imp_inverse_less linorder_neqE_ordered_idom nz positive_imp_inverse_positive real_le_antisym real_less_def w) |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
77 |
hence "inverse(real(Suc(natceiling(inverse(g w - f w))))) < g w - f w" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
78 |
by (metis inverse_inverse_eq order_less_le_trans real_le_refl) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
79 |
thus "\<exists>n. f w \<le> g w - inverse(real(Suc n))" using w |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
80 |
by (rule_tac x="natceiling(inverse(g w - f w))" in exI, auto) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
81 |
next |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
82 |
fix w n |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
83 |
assume le: "f w \<le> g w - inverse(real(Suc n))" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
84 |
hence "0 < inverse(real(Suc n))" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
85 |
by (metis inverse_real_of_nat_gt_zero) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
86 |
thus "f w < g w" using le |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
87 |
by arith |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
88 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
89 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
90 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
91 |
lemma (in sigma_algebra) sigma_le_less: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
92 |
assumes M: "!!a::real. {w \<in> space M. f w \<le> a} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
93 |
shows "{w \<in> space M. f w < a} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
94 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
95 |
show ?thesis using Collect_less_le [of "space M" f "\<lambda>x. a"] |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
96 |
by (auto simp add: countable_UN M) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
97 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
98 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
99 |
lemma (in sigma_algebra) sigma_less_ge: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
100 |
assumes M: "!!a::real. {w \<in> space M. f w < a} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
101 |
shows "{w \<in> space M. a \<le> f w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
102 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
103 |
have "{w \<in> space M. a \<le> f w} = space M - {w \<in> space M. f w < a}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
104 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
105 |
thus ?thesis using M |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
106 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
107 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
108 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
109 |
lemma (in sigma_algebra) sigma_ge_gr: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
110 |
assumes M: "!!a::real. {w \<in> space M. a \<le> f w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
111 |
shows "{w \<in> space M. a < f w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
112 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
113 |
show ?thesis using Collect_less_le [of "space M" "\<lambda>x. a" f] |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
114 |
by (auto simp add: countable_UN le_diff_eq M) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
115 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
116 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
117 |
lemma (in sigma_algebra) sigma_gr_le: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
118 |
assumes M: "!!a::real. {w \<in> space M. a < f w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
119 |
shows "{w \<in> space M. f w \<le> a} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
120 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
121 |
have "{w \<in> space M. f w \<le> a} = space M - {w \<in> space M. a < f w}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
122 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
123 |
thus ?thesis |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
124 |
by (simp add: M compl_sets) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
125 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
126 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
127 |
lemma (in measure_space) borel_measurable_gr_iff: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
128 |
"f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
129 |
proof (auto simp add: borel_measurable_le_iff sigma_gr_le) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
130 |
fix u |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
131 |
assume M: "\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
132 |
have "{w \<in> space M. u < f w} = space M - {w \<in> space M. f w \<le> u}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
133 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
134 |
thus "{w \<in> space M. u < f w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
135 |
by (force simp add: compl_sets countable_UN M) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
136 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
137 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
138 |
lemma (in measure_space) borel_measurable_less_iff: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
139 |
"f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
140 |
proof (auto simp add: borel_measurable_le_iff sigma_le_less) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
141 |
fix u |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
142 |
assume M: "\<forall>a. {w \<in> space M. f w < a} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
143 |
have "{w \<in> space M. f w \<le> u} = space M - {w \<in> space M. u < f w}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
144 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
145 |
thus "{w \<in> space M. f w \<le> u} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
146 |
using Collect_less_le [of "space M" "\<lambda>x. u" f] |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
147 |
by (force simp add: compl_sets countable_UN le_diff_eq sigma_less_ge M) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
148 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
149 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
150 |
lemma (in measure_space) borel_measurable_ge_iff: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
151 |
"f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
152 |
proof (auto simp add: borel_measurable_less_iff sigma_le_less sigma_ge_gr sigma_gr_le) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
153 |
fix u |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
154 |
assume M: "\<forall>a. {w \<in> space M. f w < a} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
155 |
have "{w \<in> space M. u \<le> f w} = space M - {w \<in> space M. f w < u}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
156 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
157 |
thus "{w \<in> space M. u \<le> f w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
158 |
by (force simp add: compl_sets countable_UN M) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
159 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
160 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
161 |
lemma (in measure_space) affine_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
162 |
assumes g: "g \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
163 |
shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
164 |
proof (cases rule: linorder_cases [of b 0]) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
165 |
case equal thus ?thesis |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
166 |
by (simp add: borel_measurable_const) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
167 |
next |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
168 |
case less |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
169 |
{ |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
170 |
fix w c |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
171 |
have "a + g w * b \<le> c \<longleftrightarrow> g w * b \<le> c - a" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
172 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
173 |
also have "... \<longleftrightarrow> (c-a)/b \<le> g w" using less |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
174 |
by (metis divide_le_eq less less_asym) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
175 |
finally have "a + g w * b \<le> c \<longleftrightarrow> (c-a)/b \<le> g w" . |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
176 |
} |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
177 |
hence "\<And>w c. a + g w * b \<le> c \<longleftrightarrow> (c-a)/b \<le> g w" . |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
178 |
thus ?thesis using less g |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
179 |
by (simp add: borel_measurable_ge_iff [of g]) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
180 |
(simp add: borel_measurable_le_iff) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
181 |
next |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
182 |
case greater |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
183 |
hence 0: "\<And>x c. (g x * b \<le> c - a) \<longleftrightarrow> (g x \<le> (c - a) / b)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
184 |
by (metis mult_imp_le_div_pos le_divide_eq) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
185 |
have 1: "\<And>x c. (a + g x * b \<le> c) \<longleftrightarrow> (g x * b \<le> c - a)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
186 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
187 |
from greater g |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
188 |
show ?thesis |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
189 |
by (simp add: borel_measurable_le_iff 0 1) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
190 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
191 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
192 |
definition |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
193 |
nat_to_rat_surj :: "nat \<Rightarrow> rat" where |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
194 |
"nat_to_rat_surj n = (let (i,j) = nat_to_nat2 n |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
195 |
in Fract (nat_to_int_bij i) (nat_to_int_bij j))" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
196 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
197 |
lemma nat_to_rat_surj: "surj nat_to_rat_surj" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
198 |
proof (auto simp add: surj_def nat_to_rat_surj_def) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
199 |
fix y |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
200 |
show "\<exists>x. y = (\<lambda>(i, j). Fract (nat_to_int_bij i) (nat_to_int_bij j)) (nat_to_nat2 x)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
201 |
proof (cases y) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
202 |
case (Fract a b) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
203 |
obtain i where i: "nat_to_int_bij i = a" using surj_nat_to_int_bij |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
204 |
by (metis surj_def) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
205 |
obtain j where j: "nat_to_int_bij j = b" using surj_nat_to_int_bij |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
206 |
by (metis surj_def) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
207 |
obtain n where n: "nat_to_nat2 n = (i,j)" using nat_to_nat2_surj |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
208 |
by (metis surj_def) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
209 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
210 |
from Fract i j n show ?thesis |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
211 |
by (metis prod.cases prod_case_split) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
212 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
213 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
214 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
215 |
lemma rats_enumeration: "\<rat> = range (of_rat o nat_to_rat_surj)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
216 |
using nat_to_rat_surj |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
217 |
by (auto simp add: image_def surj_def) (metis Rats_cases) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
218 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
219 |
lemma (in measure_space) borel_measurable_less_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
220 |
assumes f: "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
221 |
assumes g: "g \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
222 |
shows "{w \<in> space M. f w < g w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
223 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
224 |
have "{w \<in> space M. f w < g w} = |
33536 | 225 |
(\<Union>r\<in>\<rat>. {w \<in> space M. f w < r} \<inter> {w \<in> space M. r < g w })" |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
226 |
by (auto simp add: Rats_dense_in_real) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
227 |
thus ?thesis using f g |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
228 |
by (simp add: borel_measurable_less_iff [of f] |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
229 |
borel_measurable_gr_iff [of g]) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
230 |
(blast intro: gen_countable_UN [OF rats_enumeration]) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
231 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
232 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
233 |
lemma (in measure_space) borel_measurable_leq_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
234 |
assumes f: "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
235 |
assumes g: "g \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
236 |
shows "{w \<in> space M. f w \<le> g w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
237 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
238 |
have "{w \<in> space M. f w \<le> g w} = space M - {w \<in> space M. g w < f w}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
239 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
240 |
thus ?thesis using f g |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
241 |
by (simp add: borel_measurable_less_borel_measurable compl_sets) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
242 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
243 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
244 |
lemma (in measure_space) borel_measurable_eq_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
245 |
assumes f: "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
246 |
assumes g: "g \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
247 |
shows "{w \<in> space M. f w = g w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
248 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
249 |
have "{w \<in> space M. f w = g w} = |
33536 | 250 |
{w \<in> space M. f w \<le> g w} \<inter> {w \<in> space M. g w \<le> f w}" |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
251 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
252 |
thus ?thesis using f g |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
253 |
by (simp add: borel_measurable_leq_borel_measurable Int) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
254 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
255 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
256 |
lemma (in measure_space) borel_measurable_neq_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
257 |
assumes f: "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
258 |
assumes g: "g \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
259 |
shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
260 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
261 |
have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
262 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
263 |
thus ?thesis using f g |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
264 |
by (simp add: borel_measurable_eq_borel_measurable compl_sets) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
265 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
266 |
|
33535 | 267 |
lemma (in measure_space) borel_measurable_add_borel_measurable: |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
268 |
assumes f: "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
269 |
assumes g: "g \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
270 |
shows "(\<lambda>x. f x + g x) \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
271 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
272 |
have 1:"!!a. {w \<in> space M. a \<le> f w + g w} = {w \<in> space M. a + (g w) * -1 \<le> f w}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
273 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
274 |
have "!!a. (\<lambda>w. a + (g w) * -1) \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
275 |
by (rule affine_borel_measurable [OF g]) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
276 |
hence "!!a. {w \<in> space M. (\<lambda>w. a + (g w) * -1)(w) \<le> f w} \<in> sets M" using f |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
277 |
by (rule borel_measurable_leq_borel_measurable) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
278 |
hence "!!a. {w \<in> space M. a \<le> f w + g w} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
279 |
by (simp add: 1) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
280 |
thus ?thesis |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
281 |
by (simp add: borel_measurable_ge_iff) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
282 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
283 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
284 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
285 |
lemma (in measure_space) borel_measurable_square: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
286 |
assumes f: "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
287 |
shows "(\<lambda>x. (f x)^2) \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
288 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
289 |
{ |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
290 |
fix a |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
291 |
have "{w \<in> space M. (f w)\<twosuperior> \<le> a} \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
292 |
proof (cases rule: linorder_cases [of a 0]) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
293 |
case less |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
294 |
hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} = {}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
295 |
by auto (metis less order_le_less_trans power2_less_0) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
296 |
also have "... \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
297 |
by (rule empty_sets) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
298 |
finally show ?thesis . |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
299 |
next |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
300 |
case equal |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
301 |
hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} = |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
302 |
{w \<in> space M. f w \<le> 0} \<inter> {w \<in> space M. 0 \<le> f w}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
303 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
304 |
also have "... \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
305 |
apply (insert f) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
306 |
apply (rule Int) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
307 |
apply (simp add: borel_measurable_le_iff) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
308 |
apply (simp add: borel_measurable_ge_iff) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
309 |
done |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
310 |
finally show ?thesis . |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
311 |
next |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
312 |
case greater |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
313 |
have "\<forall>x. (f x ^ 2 \<le> sqrt a ^ 2) = (- sqrt a \<le> f x & f x \<le> sqrt a)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
314 |
by (metis abs_le_interval_iff abs_of_pos greater real_sqrt_abs |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
315 |
real_sqrt_le_iff real_sqrt_power) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
316 |
hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} = |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
317 |
{w \<in> space M. -(sqrt a) \<le> f w} \<inter> {w \<in> space M. f w \<le> sqrt a}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
318 |
using greater by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
319 |
also have "... \<in> sets M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
320 |
apply (insert f) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
321 |
apply (rule Int) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
322 |
apply (simp add: borel_measurable_ge_iff) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
323 |
apply (simp add: borel_measurable_le_iff) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
324 |
done |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
325 |
finally show ?thesis . |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
326 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
327 |
} |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
328 |
thus ?thesis by (auto simp add: borel_measurable_le_iff) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
329 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
330 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
331 |
lemma times_eq_sum_squares: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
332 |
fixes x::real |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
333 |
shows"x*y = ((x+y)^2)/4 - ((x-y)^ 2)/4" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
334 |
by (simp add: power2_eq_square ring_distribs diff_divide_distrib [symmetric]) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
335 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
336 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
337 |
lemma (in measure_space) borel_measurable_uminus_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
338 |
assumes g: "g \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
339 |
shows "(\<lambda>x. - g x) \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
340 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
341 |
have "(\<lambda>x. - g x) = (\<lambda>x. 0 + (g x) * -1)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
342 |
by simp |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
343 |
also have "... \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
344 |
by (fast intro: affine_borel_measurable g) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
345 |
finally show ?thesis . |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
346 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
347 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
348 |
lemma (in measure_space) borel_measurable_times_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
349 |
assumes f: "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
350 |
assumes g: "g \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
351 |
shows "(\<lambda>x. f x * g x) \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
352 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
353 |
have 1: "(\<lambda>x. 0 + (f x + g x)\<twosuperior> * inverse 4) \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
354 |
by (fast intro: affine_borel_measurable borel_measurable_square |
33535 | 355 |
borel_measurable_add_borel_measurable f g) |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
356 |
have "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) = |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
357 |
(\<lambda>x. 0 + ((f x + -g x) ^ 2 * inverse -4))" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
358 |
by (simp add: Ring_and_Field.minus_divide_right) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
359 |
also have "... \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
360 |
by (fast intro: affine_borel_measurable borel_measurable_square |
33535 | 361 |
borel_measurable_add_borel_measurable |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
362 |
borel_measurable_uminus_borel_measurable f g) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
363 |
finally have 2: "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) \<in> borel_measurable M" . |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
364 |
show ?thesis |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
365 |
apply (simp add: times_eq_sum_squares real_diff_def) |
33535 | 366 |
using 1 2 apply (simp add: borel_measurable_add_borel_measurable) |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
367 |
done |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
368 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
369 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
370 |
lemma (in measure_space) borel_measurable_diff_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
371 |
assumes f: "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
372 |
assumes g: "g \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
373 |
shows "(\<lambda>x. f x - g x) \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
374 |
unfolding real_diff_def |
33535 | 375 |
by (fast intro: borel_measurable_add_borel_measurable |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
376 |
borel_measurable_uminus_borel_measurable f g) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
377 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
378 |
lemma (in measure_space) mono_convergent_borel_measurable: |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
379 |
assumes u: "!!n. u n \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
380 |
assumes mc: "mono_convergent u f (space M)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
381 |
shows "f \<in> borel_measurable M" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
382 |
proof - |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
383 |
{ |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
384 |
fix a |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
385 |
have 1: "!!w. w \<in> space M & f w <= a \<longleftrightarrow> w \<in> space M & (\<forall>i. u i w <= a)" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
386 |
proof safe |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
387 |
fix w i |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
388 |
assume w: "w \<in> space M" and f: "f w \<le> a" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
389 |
hence "u i w \<le> f w" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
390 |
by (auto intro: SEQ.incseq_le |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
391 |
simp add: incseq_def mc [unfolded mono_convergent_def]) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
392 |
thus "u i w \<le> a" using f |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
393 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
394 |
next |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
395 |
fix w |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
396 |
assume w: "w \<in> space M" and u: "\<forall>i. u i w \<le> a" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
397 |
thus "f w \<le> a" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
398 |
by (metis LIMSEQ_le_const2 mc [unfolded mono_convergent_def]) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
399 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
400 |
have "{w \<in> space M. f w \<le> a} = {w \<in> space M. (\<forall>i. u i w <= a)}" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
401 |
by (simp add: 1) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
402 |
also have "... = (\<Inter>i. {w \<in> space M. u i w \<le> a})" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
403 |
by auto |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
404 |
also have "... \<in> sets M" using u |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
405 |
by (auto simp add: borel_measurable_le_iff intro: countable_INT) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
406 |
finally have "{w \<in> space M. f w \<le> a} \<in> sets M" . |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
407 |
} |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
408 |
thus ?thesis |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
409 |
by (auto simp add: borel_measurable_le_iff) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
410 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
411 |
|
33535 | 412 |
lemma (in measure_space) borel_measurable_setsum_borel_measurable: |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
413 |
assumes s: "finite s" |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
414 |
shows "(!!i. i \<in> s ==> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) s) \<in> borel_measurable M" using s |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
415 |
proof (induct s) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
416 |
case empty |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
417 |
thus ?case |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
418 |
by (simp add: borel_measurable_const) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
419 |
next |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
420 |
case (insert x s) |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
421 |
thus ?case |
33535 | 422 |
by (auto simp add: borel_measurable_add_borel_measurable) |
33533
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
423 |
qed |
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
424 |
|
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
paulson
parents:
diff
changeset
|
425 |
end |