src/HOL/Library/Indicator_Function.thy
changeset 57446 06e195515deb
parent 57418 6ab1c7cb0b8d
child 57447 87429bdecad5
--- a/src/HOL/Library/Indicator_Function.thy	Tue Jul 01 11:06:31 2014 +0200
+++ b/src/HOL/Library/Indicator_Function.thy	Mon Jun 30 15:45:03 2014 +0200
@@ -28,20 +28,25 @@
 lemma indicator_eq_1_iff: "indicator A x = (1::_::zero_neq_one) \<longleftrightarrow> x \<in> A"
   by (auto simp: indicator_def)
 
-lemma split_indicator:
-  "P (indicator S x) \<longleftrightarrow> ((x \<in> S \<longrightarrow> P 1) \<and> (x \<notin> S \<longrightarrow> P 0))"
+lemma split_indicator: "P (indicator S x) \<longleftrightarrow> ((x \<in> S \<longrightarrow> P 1) \<and> (x \<notin> S \<longrightarrow> P 0))"
+  unfolding indicator_def by auto
+
+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))"
   unfolding indicator_def by auto
 
 lemma indicator_inter_arith: "indicator (A \<inter> B) x = indicator A x * (indicator B x::'a::semiring_1)"
   unfolding indicator_def by (auto simp: min_def max_def)
 
-lemma indicator_union_arith: "indicator (A \<union> B) x =  indicator A x + indicator B x - indicator A x * (indicator B x::'a::ring_1)"
+lemma indicator_union_arith: "indicator (A \<union> B) x = indicator A x + indicator B x - indicator A x * (indicator B x::'a::ring_1)"
   unfolding indicator_def by (auto simp: min_def max_def)
 
 lemma indicator_inter_min: "indicator (A \<inter> B) x = min (indicator A x) (indicator B x::'a::linordered_semidom)"
   and indicator_union_max: "indicator (A \<union> B) x = max (indicator A x) (indicator B x::'a::linordered_semidom)"
   unfolding indicator_def by (auto simp: min_def max_def)
 
+lemma indicator_disj_union: "A \<inter> B = {} \<Longrightarrow>  indicator (A \<union> B) x = (indicator A x + indicator B x::'a::linordered_semidom)"
+  by (auto split: split_indicator)
+
 lemma indicator_compl: "indicator (- A) x = 1 - (indicator A x::'a::ring_1)"
   and indicator_diff: "indicator (A - B) x = indicator A x * (1 - indicator B x::'a::ring_1)"
   unfolding indicator_def by (auto simp: min_def max_def)
@@ -70,4 +75,71 @@
     (\<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)"
   using assms by (auto intro!: setsum.mono_neutral_cong_right split: split_if_asm simp: indicator_def)
 
-end
\ No newline at end of file
+lemma LIMSEQ_indicator_incseq:
+  assumes "incseq A"
+  shows "(\<lambda>i. indicator (A i) x :: 'a :: {topological_space, one, zero}) ----> indicator (\<Union>i. A i) x"
+proof cases
+  assume "\<exists>i. x \<in> A i"
+  then obtain i where "x \<in> A i"
+    by auto
+  then have 
+    "\<And>n. (indicator (A (n + i)) x :: 'a) = 1"
+    "(indicator (\<Union> i. A i) x :: 'a) = 1"
+    using incseqD[OF `incseq A`, of i "n + i" for n] `x \<in> A i` by (auto simp: indicator_def)
+  then show ?thesis
+    by (rule_tac LIMSEQ_offset[of _ i]) (simp add: tendsto_const)
+qed (auto simp: indicator_def tendsto_const)
+
+lemma LIMSEQ_indicator_UN:
+  "(\<lambda>k. indicator (\<Union> i<k. A i) x :: 'a :: {topological_space, one, zero}) ----> indicator (\<Union>i. A i) x"
+proof -
+  have "(\<lambda>k. indicator (\<Union> i<k. A i) x::'a) ----> indicator (\<Union>k. \<Union> i<k. A i) x"
+    by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def intro: less_le_trans)
+  also have "(\<Union>k. \<Union> i<k. A i) = (\<Union>i. A i)"
+    by auto
+  finally show ?thesis .
+qed
+
+lemma LIMSEQ_indicator_decseq:
+  assumes "decseq A"
+  shows "(\<lambda>i. indicator (A i) x :: 'a :: {topological_space, one, zero}) ----> indicator (\<Inter>i. A i) x"
+proof cases
+  assume "\<exists>i. x \<notin> A i"
+  then obtain i where "x \<notin> A i"
+    by auto
+  then have 
+    "\<And>n. (indicator (A (n + i)) x :: 'a) = 0"
+    "(indicator (\<Inter>i. A i) x :: 'a) = 0"
+    using decseqD[OF `decseq A`, of i "n + i" for n] `x \<notin> A i` by (auto simp: indicator_def)
+  then show ?thesis
+    by (rule_tac LIMSEQ_offset[of _ i]) (simp add: tendsto_const)
+qed (auto simp: indicator_def tendsto_const)
+
+lemma LIMSEQ_indicator_INT:
+  "(\<lambda>k. indicator (\<Inter>i<k. A i) x :: 'a :: {topological_space, one, zero}) ----> indicator (\<Inter>i. A i) x"
+proof -
+  have "(\<lambda>k. indicator (\<Inter>i<k. A i) x::'a) ----> indicator (\<Inter>k. \<Inter>i<k. A i) x"
+    by (intro LIMSEQ_indicator_decseq) (auto simp: decseq_def intro: less_le_trans)
+  also have "(\<Inter>k. \<Inter> i<k. A i) = (\<Inter>i. A i)"
+    by auto
+  finally show ?thesis .
+qed
+
+lemma indicator_add:
+  "A \<inter> B = {} \<Longrightarrow> (indicator A x::_::monoid_add) + indicator B x = indicator (A \<union> B) x"
+  unfolding indicator_def by auto
+
+lemma of_real_indicator: "of_real (indicator A x) = indicator A x"
+  by (simp split: split_indicator)
+
+lemma real_of_nat_indicator: "real (indicator A x :: nat) = indicator A x"
+  by (simp split: split_indicator)
+
+lemma abs_indicator: "\<bar>indicator A x :: 'a::linordered_idom\<bar> = indicator A x"
+  by (simp split: split_indicator)
+
+lemma mult_indicator_subset:
+  "A \<subseteq> B \<Longrightarrow> indicator A x * indicator B x = (indicator A x :: 'a::{comm_semiring_1})"
+  by (auto split: split_indicator simp: fun_eq_iff)
+
+end