src/HOL/Library/Indicator_Function.thy
changeset 54408 67dec4ccaabd
parent 45425 7fee7d7abf2f
child 56993 e5366291d6aa
--- a/src/HOL/Library/Indicator_Function.thy	Tue Nov 12 14:24:34 2013 +0100
+++ b/src/HOL/Library/Indicator_Function.thy	Tue Nov 12 19:28:50 2013 +0100
@@ -22,6 +22,12 @@
 lemma indicator_abs_le_1: "\<bar>indicator S x\<bar> \<le> (1::'a::linordered_idom)"
   unfolding indicator_def by auto
 
+lemma indicator_eq_0_iff: "indicator A x = (0::_::zero_neq_one) \<longleftrightarrow> x \<notin> A"
+  by (auto simp: indicator_def)
+
+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))"
   unfolding indicator_def by auto