author | hoelzl |
Mon, 30 Jun 2014 15:45:21 +0200 | |
changeset 57447 | 87429bdecad5 |
parent 57446 | 06e195515deb |
child 58729 | e8ecc79aee43 |
permissions | -rw-r--r-- |
37665 | 1 |
(* Title: HOL/Library/Indicator_Function.thy |
2 |
Author: Johannes Hoelzl (TU Muenchen) |
|
3 |
*) |
|
4 |
||
5 |
header {* Indicator Function *} |
|
6 |
||
7 |
theory Indicator_Function |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
54408
diff
changeset
|
8 |
imports Complex_Main |
37665 | 9 |
begin |
10 |
||
11 |
definition "indicator S x = (if x \<in> S then 1 else 0)" |
|
12 |
||
13 |
lemma indicator_simps[simp]: |
|
14 |
"x \<in> S \<Longrightarrow> indicator S x = 1" |
|
15 |
"x \<notin> S \<Longrightarrow> indicator S x = 0" |
|
16 |
unfolding indicator_def by auto |
|
17 |
||
45425 | 18 |
lemma indicator_pos_le[intro, simp]: "(0::'a::linordered_semidom) \<le> indicator S x" |
37665 | 19 |
and indicator_le_1[intro, simp]: "indicator S x \<le> (1::'a::linordered_semidom)" |
45425 | 20 |
unfolding indicator_def by auto |
21 |
||
22 |
lemma indicator_abs_le_1: "\<bar>indicator S x\<bar> \<le> (1::'a::linordered_idom)" |
|
37665 | 23 |
unfolding indicator_def by auto |
24 |
||
54408 | 25 |
lemma indicator_eq_0_iff: "indicator A x = (0::_::zero_neq_one) \<longleftrightarrow> x \<notin> A" |
26 |
by (auto simp: indicator_def) |
|
27 |
||
28 |
lemma indicator_eq_1_iff: "indicator A x = (1::_::zero_neq_one) \<longleftrightarrow> x \<in> A" |
|
29 |
by (auto simp: indicator_def) |
|
30 |
||
57446
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
31 |
lemma split_indicator: "P (indicator S x) \<longleftrightarrow> ((x \<in> S \<longrightarrow> P 1) \<and> (x \<notin> S \<longrightarrow> P 0))" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
32 |
unfolding indicator_def by auto |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
33 |
|
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
34 |
lemma split_indicator_asm: "P (indicator S x) \<longleftrightarrow> (\<not> (x \<in> S \<and> \<not> P 1 \<or> x \<notin> S \<and> \<not> P 0))" |
37665 | 35 |
unfolding indicator_def by auto |
36 |
||
45425 | 37 |
lemma indicator_inter_arith: "indicator (A \<inter> B) x = indicator A x * (indicator B x::'a::semiring_1)" |
38 |
unfolding indicator_def by (auto simp: min_def max_def) |
|
39 |
||
57446
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
40 |
lemma indicator_union_arith: "indicator (A \<union> B) x = indicator A x + indicator B x - indicator A x * (indicator B x::'a::ring_1)" |
45425 | 41 |
unfolding indicator_def by (auto simp: min_def max_def) |
42 |
||
43 |
lemma indicator_inter_min: "indicator (A \<inter> B) x = min (indicator A x) (indicator B x::'a::linordered_semidom)" |
|
37665 | 44 |
and indicator_union_max: "indicator (A \<union> B) x = max (indicator A x) (indicator B x::'a::linordered_semidom)" |
45425 | 45 |
unfolding indicator_def by (auto simp: min_def max_def) |
46 |
||
57446
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
47 |
lemma indicator_disj_union: "A \<inter> B = {} \<Longrightarrow> indicator (A \<union> B) x = (indicator A x + indicator B x::'a::linordered_semidom)" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
48 |
by (auto split: split_indicator) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
49 |
|
45425 | 50 |
lemma indicator_compl: "indicator (- A) x = 1 - (indicator A x::'a::ring_1)" |
37665 | 51 |
and indicator_diff: "indicator (A - B) x = indicator A x * (1 - indicator B x::'a::ring_1)" |
52 |
unfolding indicator_def by (auto simp: min_def max_def) |
|
53 |
||
45425 | 54 |
lemma indicator_times: "indicator (A \<times> B) x = indicator A (fst x) * (indicator B (snd x)::'a::semiring_1)" |
37665 | 55 |
unfolding indicator_def by (cases x) auto |
56 |
||
45425 | 57 |
lemma indicator_sum: "indicator (A <+> B) x = (case x of Inl x \<Rightarrow> indicator A x | Inr x \<Rightarrow> indicator B x)" |
37665 | 58 |
unfolding indicator_def by (cases x) auto |
59 |
||
60 |
lemma |
|
61 |
fixes f :: "'a \<Rightarrow> 'b::semiring_1" assumes "finite A" |
|
62 |
shows setsum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator B x) = (\<Sum>x \<in> A \<inter> B. f x)" |
|
63 |
and setsum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator B x * f x) = (\<Sum>x \<in> A \<inter> B. f x)" |
|
64 |
unfolding indicator_def |
|
57418 | 65 |
using assms by (auto intro!: setsum.mono_neutral_cong_right split: split_if_asm) |
37665 | 66 |
|
67 |
lemma setsum_indicator_eq_card: |
|
68 |
assumes "finite A" |
|
69 |
shows "(SUM x : A. indicator B x) = card (A Int B)" |
|
70 |
using setsum_mult_indicator[OF assms, of "%x. 1::nat"] |
|
71 |
unfolding card_eq_setsum by simp |
|
72 |
||
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
54408
diff
changeset
|
73 |
lemma setsum_indicator_scaleR[simp]: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
54408
diff
changeset
|
74 |
"finite A \<Longrightarrow> |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
54408
diff
changeset
|
75 |
(\<Sum>x \<in> A. indicator (B x) (g x) *\<^sub>R f x) = (\<Sum>x \<in> {x\<in>A. g x \<in> B x}. f x::'a::real_vector)" |
57418 | 76 |
using assms by (auto intro!: setsum.mono_neutral_cong_right split: split_if_asm simp: indicator_def) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
54408
diff
changeset
|
77 |
|
57446
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
78 |
lemma LIMSEQ_indicator_incseq: |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
79 |
assumes "incseq A" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
80 |
shows "(\<lambda>i. indicator (A i) x :: 'a :: {topological_space, one, zero}) ----> indicator (\<Union>i. A i) x" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
81 |
proof cases |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
82 |
assume "\<exists>i. x \<in> A i" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
83 |
then obtain i where "x \<in> A i" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
84 |
by auto |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
85 |
then have |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
86 |
"\<And>n. (indicator (A (n + i)) x :: 'a) = 1" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
87 |
"(indicator (\<Union> i. A i) x :: 'a) = 1" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
88 |
using incseqD[OF `incseq A`, of i "n + i" for n] `x \<in> A i` by (auto simp: indicator_def) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
89 |
then show ?thesis |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
90 |
by (rule_tac LIMSEQ_offset[of _ i]) (simp add: tendsto_const) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
91 |
qed (auto simp: indicator_def tendsto_const) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
92 |
|
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
93 |
lemma LIMSEQ_indicator_UN: |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
94 |
"(\<lambda>k. indicator (\<Union> i<k. A i) x :: 'a :: {topological_space, one, zero}) ----> indicator (\<Union>i. A i) x" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
95 |
proof - |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
96 |
have "(\<lambda>k. indicator (\<Union> i<k. A i) x::'a) ----> indicator (\<Union>k. \<Union> i<k. A i) x" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
97 |
by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def intro: less_le_trans) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
98 |
also have "(\<Union>k. \<Union> i<k. A i) = (\<Union>i. A i)" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
99 |
by auto |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
100 |
finally show ?thesis . |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
101 |
qed |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
102 |
|
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
103 |
lemma LIMSEQ_indicator_decseq: |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
104 |
assumes "decseq A" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
105 |
shows "(\<lambda>i. indicator (A i) x :: 'a :: {topological_space, one, zero}) ----> indicator (\<Inter>i. A i) x" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
106 |
proof cases |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
107 |
assume "\<exists>i. x \<notin> A i" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
108 |
then obtain i where "x \<notin> A i" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
109 |
by auto |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
110 |
then have |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
111 |
"\<And>n. (indicator (A (n + i)) x :: 'a) = 0" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
112 |
"(indicator (\<Inter>i. A i) x :: 'a) = 0" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
113 |
using decseqD[OF `decseq A`, of i "n + i" for n] `x \<notin> A i` by (auto simp: indicator_def) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
114 |
then show ?thesis |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
115 |
by (rule_tac LIMSEQ_offset[of _ i]) (simp add: tendsto_const) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
116 |
qed (auto simp: indicator_def tendsto_const) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
117 |
|
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
118 |
lemma LIMSEQ_indicator_INT: |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
119 |
"(\<lambda>k. indicator (\<Inter>i<k. A i) x :: 'a :: {topological_space, one, zero}) ----> indicator (\<Inter>i. A i) x" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
120 |
proof - |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
121 |
have "(\<lambda>k. indicator (\<Inter>i<k. A i) x::'a) ----> indicator (\<Inter>k. \<Inter>i<k. A i) x" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
122 |
by (intro LIMSEQ_indicator_decseq) (auto simp: decseq_def intro: less_le_trans) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
123 |
also have "(\<Inter>k. \<Inter> i<k. A i) = (\<Inter>i. A i)" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
124 |
by auto |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
125 |
finally show ?thesis . |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
126 |
qed |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
127 |
|
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
128 |
lemma indicator_add: |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
129 |
"A \<inter> B = {} \<Longrightarrow> (indicator A x::_::monoid_add) + indicator B x = indicator (A \<union> B) x" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
130 |
unfolding indicator_def by auto |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
131 |
|
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
132 |
lemma of_real_indicator: "of_real (indicator A x) = indicator A x" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
133 |
by (simp split: split_indicator) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
134 |
|
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
135 |
lemma real_of_nat_indicator: "real (indicator A x :: nat) = indicator A x" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
136 |
by (simp split: split_indicator) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
137 |
|
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
138 |
lemma abs_indicator: "\<bar>indicator A x :: 'a::linordered_idom\<bar> = indicator A x" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
139 |
by (simp split: split_indicator) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
140 |
|
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
141 |
lemma mult_indicator_subset: |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
142 |
"A \<subseteq> B \<Longrightarrow> indicator A x * indicator B x = (indicator A x :: 'a::{comm_semiring_1})" |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
143 |
by (auto split: split_indicator simp: fun_eq_iff) |
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
144 |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
145 |
lemma indicator_sums: |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
146 |
assumes "\<And>i j. i \<noteq> j \<Longrightarrow> A i \<inter> A j = {}" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
147 |
shows "(\<lambda>i. indicator (A i) x::real) sums indicator (\<Union>i. A i) x" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
148 |
proof cases |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
149 |
assume "\<exists>i. x \<in> A i" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
150 |
then obtain i where i: "x \<in> A i" .. |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
151 |
with assms have "(\<lambda>i. indicator (A i) x::real) sums (\<Sum>i\<in>{i}. indicator (A i) x)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
152 |
by (intro sums_finite) (auto split: split_indicator) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
153 |
also have "(\<Sum>i\<in>{i}. indicator (A i) x) = indicator (\<Union>i. A i) x" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
154 |
using i by (auto split: split_indicator) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
155 |
finally show ?thesis . |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
156 |
qed simp |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57446
diff
changeset
|
157 |
|
57446
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents:
57418
diff
changeset
|
158 |
end |