Add theory for indicator function.
--- a/src/HOL/IsaMakefile Thu Jul 01 09:01:09 2010 +0200
+++ b/src/HOL/IsaMakefile Thu Jul 01 11:48:42 2010 +0200
@@ -406,7 +406,8 @@
Library/Float.thy Library/Formal_Power_Series.thy \
Library/Fraction_Field.thy Library/FrechetDeriv.thy Library/Fset.thy \
Library/FuncSet.thy Library/Fundamental_Theorem_Algebra.thy \
- Library/Glbs.thy Library/Infinite_Set.thy Library/Inner_Product.thy \
+ Library/Glbs.thy Library/Indicator_Function.thy \
+ Library/Infinite_Set.thy Library/Inner_Product.thy \
Library/HOL_Library_ROOT.ML Library/Kleene_Algebra.thy \
Library/LaTeXsugar.thy Library/Lattice_Algebras.thy \
Library/Lattice_Syntax.thy Library/Library.thy \
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Indicator_Function.thy Thu Jul 01 11:48:42 2010 +0200
@@ -0,0 +1,58 @@
+(* Title: HOL/Library/Indicator_Function.thy
+ Author: Johannes Hoelzl (TU Muenchen)
+*)
+
+header {* Indicator Function *}
+
+theory Indicator_Function
+imports Main
+begin
+
+definition "indicator S x = (if x \<in> S then 1 else 0)"
+
+lemma indicator_simps[simp]:
+ "x \<in> S \<Longrightarrow> indicator S x = 1"
+ "x \<notin> S \<Longrightarrow> indicator S x = 0"
+ unfolding indicator_def by auto
+
+lemma
+ shows indicator_pos_le[intro, simp]: "(0::'a::linordered_semidom) \<le> indicator S x"
+ and indicator_le_1[intro, simp]: "indicator S x \<le> (1::'a::linordered_semidom)"
+ and indicator_abs_le_1: "\<bar>indicator S x\<bar> \<le> (1::'a::linordered_idom)"
+ unfolding indicator_def by auto
+
+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
+ shows indicator_inter_arith: "indicator (A \<inter> B) x = indicator A x * (indicator B x::'a::semiring_1)"
+ and indicator_union_arith: "indicator (A \<union> B) x = indicator A x + indicator B x - indicator A x * (indicator B x::'a::ring_1)"
+ and 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)"
+ and 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)
+
+lemma
+ shows indicator_times: "indicator (A \<times> B) x = indicator A (fst x) * (indicator B (snd x)::'a::semiring_1)"
+ unfolding indicator_def by (cases x) auto
+
+lemma
+ shows indicator_sum: "indicator (A <+> B) x = (case x of Inl x \<Rightarrow> indicator A x | Inr x \<Rightarrow> indicator B x)"
+ unfolding indicator_def by (cases x) auto
+
+lemma
+ fixes f :: "'a \<Rightarrow> 'b::semiring_1" assumes "finite A"
+ shows setsum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator B x) = (\<Sum>x \<in> A \<inter> B. f x)"
+ and setsum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator B x * f x) = (\<Sum>x \<in> A \<inter> B. f x)"
+ unfolding indicator_def
+ using assms by (auto intro!: setsum_mono_zero_cong_right split: split_if_asm)
+
+lemma setsum_indicator_eq_card:
+ assumes "finite A"
+ shows "(SUM x : A. indicator B x) = card (A Int B)"
+ using setsum_mult_indicator[OF assms, of "%x. 1::nat"]
+ unfolding card_eq_setsum by simp
+
+end
\ No newline at end of file
--- a/src/HOL/Library/Library.thy Thu Jul 01 09:01:09 2010 +0200
+++ b/src/HOL/Library/Library.thy Thu Jul 01 11:48:42 2010 +0200
@@ -27,6 +27,7 @@
Fset
FuncSet
Fundamental_Theorem_Algebra
+ Indicator_Function
Infinite_Set
Inner_Product
Lattice_Algebras
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Thu Jul 01 09:01:09 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Thu Jul 01 11:48:42 2010 +0200
@@ -179,9 +179,6 @@
subsection{* Basic componentwise operations on vectors. *}
-lemma dimindex_ge_1:"CARD(_::finite) \<ge> 1"
- using one_le_card_finite by auto
-
instantiation cart :: (times,finite) times
begin
definition vector_mult_def : "op * \<equiv> (\<lambda> x y. (\<chi> i. (x$i) * (y$i)))"
--- a/src/HOL/Multivariate_Analysis/Integration.thy Thu Jul 01 09:01:09 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Thu Jul 01 11:48:42 2010 +0200
@@ -4,7 +4,7 @@
Translation from HOL light: Robert Himmelmann, TU Muenchen *)
theory Integration
- imports Derivative "~~/src/HOL/Decision_Procs/Dense_Linear_Order"
+ imports Derivative "~~/src/HOL/Decision_Procs/Dense_Linear_Order" Indicator_Function
begin
declare [[smt_certificates="~~/src/HOL/Multivariate_Analysis/Integration.certs"]]
@@ -2468,22 +2468,7 @@
subsection {* Negligible sets. *}
-definition "indicator s \<equiv> (\<lambda>x. if x \<in> s then 1 else (0::real))"
-
-lemma dest_vec1_indicator:
- "indicator s x = (if x \<in> s then 1 else 0)" unfolding indicator_def by auto
-
-lemma indicator_pos_le[intro]: "0 \<le> (indicator s x)" unfolding indicator_def by auto
-
-lemma indicator_le_1[intro]: "(indicator s x) \<le> 1" unfolding indicator_def by auto
-
-lemma dest_vec1_indicator_abs_le_1: "abs(indicator s x) \<le> 1"
- unfolding indicator_def by auto
-
-definition "negligible (s::(_::ordered_euclidean_space) set) \<equiv> (\<forall>a b. ((indicator s) has_integral 0) {a..b})"
-
-lemma indicator_simps[simp]:"x\<in>s \<Longrightarrow> indicator s x = 1" "x\<notin>s \<Longrightarrow> indicator s x = 0"
- unfolding indicator_def by auto
+definition "negligible (s::('a::ordered_euclidean_space) set) \<equiv> (\<forall>a b. ((indicator s :: 'a\<Rightarrow>real) has_integral 0) {a..b})"
subsection {* Negligibility of hyperplane. *}
@@ -2543,7 +2528,9 @@
lemma negligible_standard_hyperplane[intro]: fixes type::"'a::ordered_euclidean_space" assumes k:"k<DIM('a)"
shows "negligible {x::'a. x$$k = (c::real)}"
unfolding negligible_def has_integral apply(rule,rule,rule,rule)
-proof- case goal1 from content_doublesplit[OF this k,of a b c] guess d . note d=this let ?i = "indicator {x::'a. x$$k = c}"
+proof-
+ case goal1 from content_doublesplit[OF this k,of a b c] guess d . note d=this
+ let ?i = "indicator {x::'a. x$$k = c} :: 'a\<Rightarrow>real"
show ?case apply(rule_tac x="\<lambda>x. ball x d" in exI) apply(rule,rule gauge_ball,rule d)
proof(rule,rule) fix p assume p:"p tagged_division_of {a..b} \<and> (\<lambda>x. ball x d) fine p"
have *:"(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. abs(x$$k - c) \<le> d}) *\<^sub>R ?i x)"
@@ -2684,13 +2671,13 @@
have "\<forall>i. \<exists>q. q tagged_division_of {a..b} \<and> (d i) fine q \<and> (\<forall>(x, k)\<in>p. k \<subseteq> (d i) x \<longrightarrow> (x, k) \<in> q)"
apply(rule,rule tagged_division_finer[OF as(1) d(1)]) by auto
from choice[OF this] guess q .. note q=conjunctD3[OF this[rule_format]]
- have *:"\<And>i. (\<Sum>(x, k)\<in>q i. content k *\<^sub>R indicator s x) \<ge> 0" apply(rule setsum_nonneg,safe)
+ have *:"\<And>i. (\<Sum>(x, k)\<in>q i. content k *\<^sub>R indicator s x) \<ge> (0::real)" apply(rule setsum_nonneg,safe)
unfolding real_scaleR_def apply(rule mult_nonneg_nonneg) apply(drule tagged_division_ofD(4)[OF q(1)]) by auto
have **:"\<And>f g s t. finite s \<Longrightarrow> finite t \<Longrightarrow> (\<forall>(x,y) \<in> t. (0::real) \<le> g(x,y)) \<Longrightarrow> (\<forall>y\<in>s. \<exists>x. (x,y) \<in> t \<and> f(y) \<le> g(x,y)) \<Longrightarrow> setsum f s \<le> setsum g t"
proof- case goal1 thus ?case apply-apply(rule setsum_le_included[of s t g snd f]) prefer 4
apply safe apply(erule_tac x=x in ballE) apply(erule exE) apply(rule_tac x="(xa,x)" in bexI) by auto qed
have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) \<le> setsum (\<lambda>i. (real i + 1) *
- norm(setsum (\<lambda>(x,k). content k *\<^sub>R indicator s x) (q i))) {0..N+1}"
+ norm(setsum (\<lambda>(x,k). content k *\<^sub>R indicator s x :: real) (q i))) {0..N+1}"
unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right
apply(rule order_trans,rule setsum_norm) defer apply(subst sum_sum_product) prefer 3
proof(rule **,safe) show "finite {(i, j) |i j. i \<in> {0..N + 1} \<and> j \<in> q i}" apply(rule finite_product_dependent) using q by auto
@@ -2780,14 +2767,14 @@
lemma negligible_unions[intro]: assumes "finite s" "\<forall>t\<in>s. negligible t" shows "negligible(\<Union>s)"
using assms by(induct,auto)
-lemma negligible: "negligible s \<longleftrightarrow> (\<forall>t::('a::ordered_euclidean_space) set. (indicator s has_integral 0) t)"
+lemma negligible: "negligible s \<longleftrightarrow> (\<forall>t::('a::ordered_euclidean_space) set. ((indicator s::'a\<Rightarrow>real) has_integral 0) t)"
apply safe defer apply(subst negligible_def)
proof- fix t::"'a set" assume as:"negligible s"
have *:"(\<lambda>x. if x \<in> s \<inter> t then 1 else 0) = (\<lambda>x. if x\<in>t then if x\<in>s then 1 else 0 else 0)" by(rule ext,auto)
- show "(indicator s has_integral 0) t" apply(subst has_integral_alt)
+ show "((indicator s::'a\<Rightarrow>real) has_integral 0) t" apply(subst has_integral_alt)
apply(cases,subst if_P,assumption) unfolding if_not_P apply(safe,rule as[unfolded negligible_def,rule_format])
apply(rule_tac x=1 in exI) apply(safe,rule zero_less_one) apply(rule_tac x=0 in exI)
- using negligible_subset[OF as,of "s \<inter> t"] unfolding negligible_def indicator_def unfolding * by auto qed auto
+ using negligible_subset[OF as,of "s \<inter> t"] unfolding negligible_def indicator_def_raw unfolding * by auto qed auto
subsection {* Finite case of the spike theorem is quite commonly needed. *}